src/HOL/Library/Bit_Operations.thy
author haftmann
Sat, 11 Jul 2020 06:21:04 +0000
changeset 72010 a851ce626b78
parent 72009 febdd4eead56
child 72023 08348e364739
permissions -rw-r--r--
signed_take_bit
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     1
(*  Author:  Florian Haftmann, TUM
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     2
*)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     3
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
     4
section \<open>Bit operations in suitable algebraic structures\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     5
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     6
theory Bit_Operations
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     7
  imports
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     8
    "HOL-Library.Boolean_Algebra"
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
     9
    Main
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    10
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    11
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    12
subsection \<open>Preliminiaries\<close> \<comment> \<open>TODO move\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    13
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    14
lemma take_bit_int_less_exp:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    15
  \<open>take_bit n k < 2 ^ n\<close> for k :: int
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    16
  by (simp add: take_bit_eq_mod)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    17
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    18
lemma take_bit_Suc_from_most:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    19
  \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> for k :: int
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    20
  by (simp only: take_bit_eq_mod power_Suc2) (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one zmod_zmult2_eq)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    21
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    22
lemma take_bit_greater_eq:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    23
  \<open>k + 2 ^ n \<le> take_bit n k\<close> if \<open>k < 0\<close> for k :: int
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    24
proof -
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    25
  have \<open>k + 2 ^ n \<le> take_bit n (k + 2 ^ n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    26
  proof (cases \<open>k > - (2 ^ n)\<close>)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    27
    case False
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    28
    then have \<open>k + 2 ^ n \<le> 0\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    29
      by simp
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    30
    also note take_bit_nonnegative
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    31
    finally show ?thesis .
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    32
  next
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    33
    case True
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    34
    with that have \<open>0 \<le> k + 2 ^ n\<close> and \<open>k + 2 ^ n < 2 ^ n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    35
      by simp_all
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    36
    then show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    37
      by (simp only: take_bit_eq_mod mod_pos_pos_trivial)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    38
  qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    39
  then show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    40
    by (simp add: take_bit_eq_mod)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    41
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    42
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    43
lemma take_bit_less_eq:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    44
  \<open>take_bit n k \<le> k - 2 ^ n\<close> if \<open>2 ^ n \<le> k\<close> and \<open>n > 0\<close> for k :: int
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    45
  using that zmod_le_nonneg_dividend [of \<open>k - 2 ^ n\<close> \<open>2 ^ n\<close>]
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    46
  by (simp add: take_bit_eq_mod)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    47
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
    48
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
    49
subsection \<open>Bit operations\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    50
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    51
class semiring_bit_operations = semiring_bit_shifts +
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    52
  fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>AND\<close> 64)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    53
    and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>OR\<close>  59)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    54
    and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>XOR\<close> 59)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    55
  assumes bit_and_iff: \<open>\<And>n. bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    56
    and bit_or_iff: \<open>\<And>n. bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    57
    and bit_xor_iff: \<open>\<And>n. bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    58
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    59
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    60
text \<open>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    61
  We want the bitwise operations to bind slightly weaker
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71042
diff changeset
    62
  than \<open>+\<close> and \<open>-\<close>.
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    63
  For the sake of code generation
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    64
  the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    65
  are specified as definitional class operations.
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    66
\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    67
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    68
sublocale "and": semilattice \<open>(AND)\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    69
  by standard (auto simp add: bit_eq_iff bit_and_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    70
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    71
sublocale or: semilattice_neutr \<open>(OR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    72
  by standard (auto simp add: bit_eq_iff bit_or_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    73
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    74
sublocale xor: comm_monoid \<open>(XOR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    75
  by standard (auto simp add: bit_eq_iff bit_xor_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    76
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    77
lemma even_and_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    78
  \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    79
  using bit_and_iff [of a b 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    80
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    81
lemma even_or_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    82
  \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    83
  using bit_or_iff [of a b 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    84
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    85
lemma even_xor_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    86
  \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    87
  using bit_xor_iff [of a b 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    88
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    89
lemma zero_and_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    90
  "0 AND a = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    91
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    92
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    93
lemma and_zero_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    94
  "a AND 0 = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    95
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    96
71921
a238074c5a9d avoid overaggressive default simp rules
haftmann
parents: 71823
diff changeset
    97
lemma one_and_eq:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    98
  "1 AND a = a mod 2"
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    99
  by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   100
71921
a238074c5a9d avoid overaggressive default simp rules
haftmann
parents: 71823
diff changeset
   101
lemma and_one_eq:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   102
  "a AND 1 = a mod 2"
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   103
  using one_and_eq [of a] by (simp add: ac_simps)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   104
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   105
lemma one_or_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   106
  "1 OR a = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   107
  by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   108
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   109
lemma or_one_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   110
  "a OR 1 = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   111
  using one_or_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   112
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   113
lemma one_xor_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   114
  "1 XOR a = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   115
  by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff odd_bit_iff_bit_pred elim: oddE)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   116
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   117
lemma xor_one_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   118
  "a XOR 1 = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   119
  using one_xor_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   120
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   121
lemma take_bit_and [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   122
  \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   123
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   124
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   125
lemma take_bit_or [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   126
  \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   127
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   128
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   129
lemma take_bit_xor [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   130
  \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   131
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   132
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   133
definition mask :: \<open>nat \<Rightarrow> 'a\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   134
  where mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   135
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   136
lemma bit_mask_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   137
  \<open>bit (mask m) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   138
  by (simp add: mask_eq_exp_minus_1 bit_mask_iff)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   139
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   140
lemma even_mask_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   141
  \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   142
  using bit_mask_iff [of n 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   143
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   144
lemma mask_0 [simp, code]:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   145
  \<open>mask 0 = 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   146
  by (simp add: mask_eq_exp_minus_1)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   147
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   148
lemma mask_Suc_exp [code]:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   149
  \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   150
  by (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   151
    (auto simp add: bit_or_iff bit_mask_iff bit_exp_iff not_less le_less_Suc_eq)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   152
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   153
lemma mask_Suc_double:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   154
  \<open>mask (Suc n) = 2 * mask n OR 1\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   155
proof (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   156
  fix q
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   157
  assume \<open>2 ^ q \<noteq> 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   158
  show \<open>bit (mask (Suc n)) q \<longleftrightarrow> bit (2 * mask n OR 1) q\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   159
    by (cases q)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   160
      (simp_all add: even_mask_iff even_or_iff bit_or_iff bit_mask_iff bit_exp_iff bit_double_iff not_less le_less_Suc_eq bit_1_iff, auto simp add: mult_2)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   161
qed
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   162
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
   163
lemma take_bit_eq_mask:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   164
  \<open>take_bit n a = a AND mask n\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   165
  by (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   166
    (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   167
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   168
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   169
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   170
class ring_bit_operations = semiring_bit_operations + ring_parity +
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   171
  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   172
  assumes bit_not_iff: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   173
  assumes minus_eq_not_minus_1: \<open>- a = NOT (a - 1)\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   174
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   175
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   176
text \<open>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   177
  For the sake of code generation \<^const>\<open>not\<close> is specified as
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   178
  definitional class operation.  Note that \<^const>\<open>not\<close> has no
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   179
  sensible definition for unlimited but only positive bit strings
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   180
  (type \<^typ>\<open>nat\<close>).
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   181
\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   182
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   183
lemma bits_minus_1_mod_2_eq [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   184
  \<open>(- 1) mod 2 = 1\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   185
  by (simp add: mod_2_eq_odd)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   186
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   187
lemma not_eq_complement:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   188
  \<open>NOT a = - a - 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   189
  using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   190
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   191
lemma minus_eq_not_plus_1:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   192
  \<open>- a = NOT a + 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   193
  using not_eq_complement [of a] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   194
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   195
lemma bit_minus_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   196
  \<open>bit (- a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit (a - 1) n\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   197
  by (simp add: minus_eq_not_minus_1 bit_not_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   198
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   199
lemma even_not_iff [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   200
  "even (NOT a) \<longleftrightarrow> odd a"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   201
  using bit_not_iff [of a 0] by auto
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   202
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   203
lemma bit_not_exp_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   204
  \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<noteq> m\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   205
  by (auto simp add: bit_not_iff bit_exp_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   206
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   207
lemma bit_minus_1_iff [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   208
  \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   209
  by (simp add: bit_minus_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   210
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   211
lemma bit_minus_exp_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   212
  \<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   213
  oops
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   214
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   215
lemma bit_minus_2_iff [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   216
  \<open>bit (- 2) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n > 0\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   217
  by (simp add: bit_minus_iff bit_1_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   218
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   219
lemma not_one [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   220
  "NOT 1 = - 2"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   221
  by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   222
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   223
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   224
  apply standard
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   225
  apply (simp add: bit_eq_iff bit_and_iff)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   226
  apply (auto simp add: exp_eq_0_imp_not_bit bit_exp_iff)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   227
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   228
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   229
sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   230
  rewrites \<open>bit.xor = (XOR)\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   231
proof -
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   232
  interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   233
    apply standard
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
   234
         apply (simp_all add: bit_eq_iff)
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
   235
       apply (auto simp add: bit_and_iff bit_or_iff bit_not_iff bit_exp_iff exp_eq_0_imp_not_bit)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   236
    done
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   237
  show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   238
    by standard
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   239
  show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
   240
    apply (simp add: fun_eq_iff bit_eq_iff bit.xor_def)
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
   241
    apply (auto simp add: bit_and_iff bit_or_iff bit_not_iff bit_xor_iff exp_eq_0_imp_not_bit)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   242
    done
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   243
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   244
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   245
lemma and_eq_not_not_or:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   246
  \<open>a AND b = NOT (NOT a OR NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   247
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   248
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   249
lemma or_eq_not_not_and:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   250
  \<open>a OR b = NOT (NOT a AND NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   251
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   252
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   253
lemma not_add_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   254
  \<open>NOT (a + b) = NOT a - b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   255
  by (simp add: not_eq_complement algebra_simps)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   256
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   257
lemma not_diff_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   258
  \<open>NOT (a - b) = NOT a + b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   259
  using not_add_distrib [of a \<open>- b\<close>] by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   260
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   261
lemma push_bit_minus:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   262
  \<open>push_bit n (- a) = - push_bit n a\<close>
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   263
  by (simp add: push_bit_eq_mult)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   264
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   265
lemma take_bit_not_take_bit:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   266
  \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   267
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   268
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   269
lemma take_bit_not_iff:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   270
  "take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   271
  apply (simp add: bit_eq_iff bit_not_iff bit_take_bit_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   272
  apply (simp add: bit_exp_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   273
  apply (use local.exp_eq_0_imp_not_bit in blast)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   274
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   275
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   276
lemma take_bit_minus_one_eq_mask:
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   277
  \<open>take_bit n (- 1) = mask n\<close>
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   278
  by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   279
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   280
lemma minus_exp_eq_not_mask:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   281
  \<open>- (2 ^ n) = NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   282
  by (rule bit_eqI) (simp add: bit_minus_iff bit_not_iff flip: mask_eq_exp_minus_1)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   283
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   284
lemma push_bit_minus_one_eq_not_mask:
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   285
  \<open>push_bit n (- 1) = NOT (mask n)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   286
  by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   287
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   288
lemma take_bit_not_mask_eq_0:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   289
  \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   290
  by (rule bit_eqI) (use that in \<open>simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\<close>)
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   291
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   292
definition set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
71991
8bff286878bf misc lemma tuning
haftmann
parents: 71986
diff changeset
   293
  where \<open>set_bit n a = a OR push_bit n 1\<close>
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   294
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   295
definition unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
71991
8bff286878bf misc lemma tuning
haftmann
parents: 71986
diff changeset
   296
  where \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   297
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   298
definition flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
71991
8bff286878bf misc lemma tuning
haftmann
parents: 71986
diff changeset
   299
  where \<open>flip_bit n a = a XOR push_bit n 1\<close>
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   300
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   301
lemma bit_set_bit_iff:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   302
  \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> 2 ^ n \<noteq> 0)\<close>
71991
8bff286878bf misc lemma tuning
haftmann
parents: 71986
diff changeset
   303
  by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   304
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   305
lemma even_set_bit_iff:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   306
  \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   307
  using bit_set_bit_iff [of m a 0] by auto
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   308
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   309
lemma bit_unset_bit_iff:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   310
  \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
71991
8bff286878bf misc lemma tuning
haftmann
parents: 71986
diff changeset
   311
  by (auto simp add: unset_bit_def push_bit_of_1 bit_and_iff bit_not_iff bit_exp_iff exp_eq_0_imp_not_bit)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   312
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   313
lemma even_unset_bit_iff:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   314
  \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   315
  using bit_unset_bit_iff [of m a 0] by auto
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   316
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   317
lemma bit_flip_bit_iff:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   318
  \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> 2 ^ n \<noteq> 0\<close>
71991
8bff286878bf misc lemma tuning
haftmann
parents: 71986
diff changeset
   319
  by (auto simp add: flip_bit_def push_bit_of_1 bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   320
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   321
lemma even_flip_bit_iff:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   322
  \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   323
  using bit_flip_bit_iff [of m a 0] by auto
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   324
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   325
lemma set_bit_0 [simp]:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   326
  \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   327
proof (rule bit_eqI)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   328
  fix m
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   329
  assume *: \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   330
  then show \<open>bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   331
    by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff)
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   332
      (cases m, simp_all add: bit_Suc)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   333
qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   334
71821
541e68d1a964 less aggressive default simp rules
haftmann
parents: 71804
diff changeset
   335
lemma set_bit_Suc:
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   336
  \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   337
proof (rule bit_eqI)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   338
  fix m
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   339
  assume *: \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   340
  show \<open>bit (set_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * set_bit n (a div 2)) m\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   341
  proof (cases m)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   342
    case 0
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   343
    then show ?thesis
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   344
      by (simp add: even_set_bit_iff)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   345
  next
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   346
    case (Suc m)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   347
    with * have \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   348
      using mult_2 by auto
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   349
    show ?thesis
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   350
      by (cases a rule: parity_cases)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   351
        (simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *,
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   352
        simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   353
  qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   354
qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   355
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   356
lemma unset_bit_0 [simp]:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   357
  \<open>unset_bit 0 a = 2 * (a div 2)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   358
proof (rule bit_eqI)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   359
  fix m
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   360
  assume *: \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   361
  then show \<open>bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   362
    by (simp add: bit_unset_bit_iff bit_double_iff)
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   363
      (cases m, simp_all add: bit_Suc)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   364
qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   365
71821
541e68d1a964 less aggressive default simp rules
haftmann
parents: 71804
diff changeset
   366
lemma unset_bit_Suc:
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   367
  \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   368
proof (rule bit_eqI)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   369
  fix m
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   370
  assume *: \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   371
  then show \<open>bit (unset_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * unset_bit n (a div 2)) m\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   372
  proof (cases m)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   373
    case 0
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   374
    then show ?thesis
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   375
      by (simp add: even_unset_bit_iff)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   376
  next
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   377
    case (Suc m)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   378
    show ?thesis
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   379
      by (cases a rule: parity_cases)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   380
        (simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *,
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   381
         simp_all add: Suc bit_Suc)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   382
  qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   383
qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   384
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   385
lemma flip_bit_0 [simp]:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   386
  \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   387
proof (rule bit_eqI)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   388
  fix m
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   389
  assume *: \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   390
  then show \<open>bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   391
    by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff)
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   392
      (cases m, simp_all add: bit_Suc)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   393
qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   394
71821
541e68d1a964 less aggressive default simp rules
haftmann
parents: 71804
diff changeset
   395
lemma flip_bit_Suc:
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   396
  \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   397
proof (rule bit_eqI)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   398
  fix m
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   399
  assume *: \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   400
  show \<open>bit (flip_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * flip_bit n (a div 2)) m\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   401
  proof (cases m)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   402
    case 0
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   403
    then show ?thesis
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   404
      by (simp add: even_flip_bit_iff)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   405
  next
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   406
    case (Suc m)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   407
    with * have \<open>2 ^ m \<noteq> 0\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   408
      using mult_2 by auto
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   409
    show ?thesis
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   410
      by (cases a rule: parity_cases)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   411
        (simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff,
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   412
        simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   413
  qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   414
qed
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   415
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   416
lemma flip_bit_eq_if:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   417
  \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   418
  by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   419
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   420
lemma take_bit_set_bit_eq:
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   421
  \<open>take_bit n (set_bit m a) = (if n \<le> m then take_bit n a else set_bit m (take_bit n a))\<close>
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   422
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   423
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   424
lemma take_bit_unset_bit_eq:
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   425
  \<open>take_bit n (unset_bit m a) = (if n \<le> m then take_bit n a else unset_bit m (take_bit n a))\<close>
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   426
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   427
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   428
lemma take_bit_flip_bit_eq:
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   429
  \<open>take_bit n (flip_bit m a) = (if n \<le> m then take_bit n a else flip_bit m (take_bit n a))\<close>
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   430
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   431
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   432
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   433
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   434
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
   435
subsection \<open>Instance \<^typ>\<open>int\<close>\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   436
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   437
instantiation int :: ring_bit_operations
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   438
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   439
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   440
definition not_int :: \<open>int \<Rightarrow> int\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   441
  where \<open>not_int k = - k - 1\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   442
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   443
lemma not_int_rec:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   444
  "NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   445
  by (auto simp add: not_int_def elim: oddE)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   446
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   447
lemma even_not_iff_int:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   448
  \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   449
  by (simp add: not_int_def)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   450
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   451
lemma not_int_div_2:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   452
  \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   453
  by (simp add: not_int_def)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   454
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   455
lemma bit_not_int_iff:
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   456
  \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   457
    for k :: int
71535
b612edee9b0c more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents: 71442
diff changeset
   458
  by (induction n arbitrary: k) (simp_all add: not_int_div_2 even_not_iff_int bit_Suc)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   459
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   460
function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   461
  where \<open>(k::int) AND l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   462
    then - of_bool (odd k \<and> odd l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   463
    else of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2)))\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   464
  by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   465
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   466
termination
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   467
  by (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>) auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   468
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   469
declare and_int.simps [simp del]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   470
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   471
lemma and_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   472
  \<open>k AND l = of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2))\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   473
    for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   474
proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   475
  case True
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   476
  then show ?thesis
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   477
    by auto (simp_all add: and_int.simps)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   478
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   479
  case False
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   480
  then show ?thesis
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   481
    by (auto simp add: ac_simps and_int.simps [of k l])
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   482
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   483
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   484
lemma bit_and_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   485
  \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   486
proof (induction n arbitrary: k l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   487
  case 0
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   488
  then show ?case
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   489
    by (simp add: and_int_rec [of k l])
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   490
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   491
  case (Suc n)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   492
  then show ?case
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   493
    by (simp add: and_int_rec [of k l] bit_Suc)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   494
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   495
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   496
lemma even_and_iff_int:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   497
  \<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   498
  using bit_and_int_iff [of k l 0] by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   499
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   500
definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   501
  where \<open>k OR l = NOT (NOT k AND NOT l)\<close> for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   502
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   503
lemma or_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   504
  \<open>k OR l = of_bool (odd k \<or> odd l) + 2 * ((k div 2) OR (l div 2))\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   505
  for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   506
  using and_int_rec [of \<open>NOT k\<close> \<open>NOT l\<close>]
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   507
  by (simp add: or_int_def even_not_iff_int not_int_div_2)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   508
    (simp add: not_int_def)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   509
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   510
lemma bit_or_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   511
  \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   512
  by (simp add: or_int_def bit_not_int_iff bit_and_int_iff)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   513
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   514
definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   515
  where \<open>k XOR l = k AND NOT l OR NOT k AND l\<close> for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   516
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   517
lemma xor_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   518
  \<open>k XOR l = of_bool (odd k \<noteq> odd l) + 2 * ((k div 2) XOR (l div 2))\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   519
  for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   520
  by (simp add: xor_int_def or_int_rec [of \<open>k AND NOT l\<close> \<open>NOT k AND l\<close>] even_and_iff_int even_not_iff_int)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   521
    (simp add: and_int_rec [of \<open>NOT k\<close> \<open>l\<close>] and_int_rec [of \<open>k\<close> \<open>NOT l\<close>] not_int_div_2)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   522
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   523
lemma bit_xor_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   524
  \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   525
  by (auto simp add: xor_int_def bit_or_int_iff bit_and_int_iff bit_not_int_iff)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   526
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   527
instance proof
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   528
  fix k l :: int and n :: nat
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   529
  show \<open>- k = NOT (k - 1)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   530
    by (simp add: not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   531
  show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   532
    by (fact bit_and_int_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   533
  show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   534
    by (fact bit_or_int_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   535
  show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   536
    by (fact bit_xor_int_iff)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   537
qed (simp_all add: bit_not_int_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   538
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   539
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   540
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   541
lemma disjunctive_add:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   542
  \<open>k + l = k OR l\<close> if \<open>\<And>n. \<not> bit k n \<or> \<not> bit l n\<close> for k l :: int
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   543
  \<comment> \<open>TODO: may integrate (indirectly) into \<^class>\<open>semiring_bits\<close> premises\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   544
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   545
  fix n
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   546
  from that have \<open>bit (k + l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   547
  proof (induction n arbitrary: k l)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   548
    case 0
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   549
    from this [of 0] show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   550
      by auto
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   551
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   552
    case (Suc n)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   553
    have \<open>bit ((k + l) div 2) n \<longleftrightarrow> bit (k div 2 + l div 2) n\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   554
      using Suc.prems [of 0] div_add1_eq [of k l] by auto
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   555
    also have \<open>bit (k div 2 + l div 2) n \<longleftrightarrow> bit (k div 2) n \<or> bit (l div 2) n\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   556
      by (rule Suc.IH) (use Suc.prems in \<open>simp flip: bit_Suc\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   557
    finally show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   558
      by (simp add: bit_Suc)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   559
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   560
  also have \<open>\<dots> \<longleftrightarrow> bit (k OR l) n\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   561
    by (simp add: bit_or_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   562
  finally show \<open>bit (k + l) n \<longleftrightarrow> bit (k OR l) n\<close> .
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   563
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   564
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   565
lemma disjunctive_diff:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   566
  \<open>k - l = k AND NOT l\<close> if \<open>\<And>n. bit l n \<Longrightarrow> bit k n\<close> for k l :: int
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   567
proof -
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   568
  have \<open>NOT k + l = NOT k OR l\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   569
    by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   570
  then have \<open>NOT (NOT k + l) = NOT (NOT k OR l)\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   571
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   572
  then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   573
    by (simp add: not_add_distrib)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   574
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   575
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   576
lemma not_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   577
  \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   578
  by (simp add: not_int_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   579
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   580
lemma not_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   581
  \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   582
  by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   583
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   584
lemma and_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   585
  \<open>k AND l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<or> l \<ge> 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   586
proof (induction k arbitrary: l rule: int_bit_induct)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   587
  case zero
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   588
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   589
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   590
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   591
  case minus
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   592
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   593
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   594
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   595
  case (even k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   596
  then show ?case
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   597
    using and_int_rec [of \<open>k * 2\<close> l] by (simp add: pos_imp_zdiv_nonneg_iff)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   598
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   599
  case (odd k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   600
  from odd have \<open>0 \<le> k AND l div 2 \<longleftrightarrow> 0 \<le> k \<or> 0 \<le> l div 2\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   601
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   602
  then have \<open>0 \<le> (1 + k * 2) div 2 AND l div 2 \<longleftrightarrow> 0 \<le> (1 + k * 2) div 2\<or> 0 \<le> l div 2\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   603
    by simp
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   604
  with and_int_rec [of \<open>1 + k * 2\<close> l]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   605
  show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   606
    by auto
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   607
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   608
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   609
lemma and_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   610
  \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   611
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   612
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   613
lemma and_less_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   614
  \<open>k AND l \<le> k\<close> if \<open>l < 0\<close> for k l :: int
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   615
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   616
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   617
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   618
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   619
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   620
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   621
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   622
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   623
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   624
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   625
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   626
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   627
    by (simp add: and_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   628
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   629
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   630
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   631
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   632
    by (simp add: and_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   633
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   634
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   635
lemma or_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   636
  \<open>k OR l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<and> l \<ge> 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   637
  by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   638
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   639
lemma or_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   640
  \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   641
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   642
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   643
lemma or_greater_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   644
  \<open>k OR l \<ge> k\<close> if \<open>l \<ge> 0\<close> for k l :: int
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   645
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   646
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   647
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   648
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   649
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   650
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   651
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   652
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   653
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   654
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   655
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   656
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   657
    by (simp add: or_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   658
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   659
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   660
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   661
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   662
    by (simp add: or_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   663
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   664
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   665
lemma xor_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   666
  \<open>k XOR l \<ge> 0 \<longleftrightarrow> (k \<ge> 0 \<longleftrightarrow> l \<ge> 0)\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   667
  by (simp only: bit.xor_def or_nonnegative_int_iff) auto
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   668
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   669
lemma xor_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   670
  \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   671
  by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   672
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   673
lemma mask_nonnegative:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   674
  \<open>(mask n :: int) \<ge> 0\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   675
 by (simp add: mask_eq_exp_minus_1)  
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   676
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   677
lemma set_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   678
  \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   679
  by (simp add: set_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   680
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   681
lemma set_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   682
  \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   683
  by (simp add: set_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   684
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   685
lemma unset_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   686
  \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   687
  by (simp add: unset_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   688
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   689
lemma unset_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   690
  \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   691
  by (simp add: unset_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   692
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   693
lemma flip_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   694
  \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   695
  by (simp add: flip_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   696
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   697
lemma flip_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   698
  \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   699
  by (simp add: flip_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   700
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   701
lemma set_bit_greater_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   702
  \<open>set_bit n k \<ge> k\<close> for k :: int
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   703
  by (simp add: set_bit_def or_greater_eq)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   704
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   705
lemma unset_bit_less_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   706
  \<open>unset_bit n k \<le> k\<close> for k :: int
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   707
  by (simp add: unset_bit_def and_less_eq)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   708
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   709
lemma set_bit_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   710
  \<open>set_bit n k = k + of_bool (\<not> bit k n) * 2 ^ n\<close> for k :: int
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   711
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   712
  fix m
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   713
  show \<open>bit (set_bit n k) m \<longleftrightarrow> bit (k + of_bool (\<not> bit k n) * 2 ^ n) m\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   714
  proof (cases \<open>m = n\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   715
    case True
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   716
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   717
      apply (simp add: bit_set_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   718
      apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   719
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   720
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   721
    case False
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   722
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   723
      apply (clarsimp simp add: bit_set_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   724
      apply (subst disjunctive_add)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   725
      apply (clarsimp simp add: bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   726
      apply (clarsimp simp add: bit_or_iff bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   727
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   728
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   729
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   730
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   731
lemma unset_bit_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   732
  \<open>unset_bit n k = k - of_bool (bit k n) * 2 ^ n\<close> for k :: int
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   733
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   734
  fix m
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   735
  show \<open>bit (unset_bit n k) m \<longleftrightarrow> bit (k - of_bool (bit k n) * 2 ^ n) m\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   736
  proof (cases \<open>m = n\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   737
    case True
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   738
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   739
      apply (simp add: bit_unset_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   740
      apply (simp add: bit_iff_odd)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   741
      using div_plus_div_distrib_dvd_right [of \<open>2 ^ n\<close> \<open>- (2 ^ n)\<close> k]
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   742
      apply (simp add: dvd_neg_div)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   743
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   744
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   745
    case False
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   746
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   747
      apply (clarsimp simp add: bit_unset_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   748
      apply (subst disjunctive_diff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   749
      apply (clarsimp simp add: bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   750
      apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   751
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   752
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   753
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   754
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
   755
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   756
subsection \<open>Taking bit with sign propagation\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   757
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   758
definition signed_take_bit :: "nat \<Rightarrow> int \<Rightarrow> int"
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   759
  where \<open>signed_take_bit n k = take_bit n k OR (NOT (mask n) * of_bool (bit k n))\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   760
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   761
lemma signed_take_bit_eq:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   762
  \<open>signed_take_bit n k = take_bit n k\<close> if \<open>\<not> bit k n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   763
  using that by (simp add: signed_take_bit_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   764
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   765
lemma signed_take_bit_eq_or:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   766
  \<open>signed_take_bit n k = take_bit n k OR NOT (mask n)\<close> if \<open>bit k n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   767
  using that by (simp add: signed_take_bit_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   768
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   769
lemma signed_take_bit_0 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   770
  \<open>signed_take_bit 0 k = - (k mod 2)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   771
  by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   772
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   773
lemma mask_half_int:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   774
  \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   775
  by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   776
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   777
lemma signed_take_bit_Suc:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   778
  \<open>signed_take_bit (Suc n) k = k mod 2 + 2 * signed_take_bit n (k div 2)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   779
  by (unfold signed_take_bit_def or_int_rec [of \<open>take_bit (Suc n) k\<close>])
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   780
    (simp add: bit_Suc take_bit_Suc even_or_iff even_mask_iff odd_iff_mod_2_eq_one not_int_div_2 mask_half_int)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   781
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   782
lemma signed_take_bit_rec:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   783
  \<open>signed_take_bit n k = (if n = 0 then - (k mod 2) else k mod 2 + 2 * signed_take_bit (n - 1) (k div 2))\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   784
  by (cases n) (simp_all add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   785
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   786
lemma bit_signed_take_bit_iff:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   787
  \<open>bit (signed_take_bit m k) n = bit k (min m n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   788
  by (simp add: signed_take_bit_def bit_or_iff bit_take_bit_iff bit_not_iff bit_mask_iff min_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   789
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   790
text \<open>Modulus centered around 0\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   791
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   792
lemma signed_take_bit_eq_take_bit_minus:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   793
  \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   794
proof (cases \<open>bit k n\<close>)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   795
  case True
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   796
  have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   797
    by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   798
  then have \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   799
    by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   800
  with True show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   801
    by (simp flip: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   802
next
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   803
  case False
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   804
  then show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   805
    by (simp add: bit_eq_iff bit_take_bit_iff bit_signed_take_bit_iff min_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   806
      (auto intro: bit_eqI simp add: less_Suc_eq)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   807
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   808
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   809
lemma signed_take_bit_eq_take_bit_shift:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   810
  \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   811
proof -
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   812
  have *: \<open>take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   813
    by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   814
  have \<open>take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   815
    by (simp add: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   816
  also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   817
    by (rule disjunctive_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   818
      (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   819
  finally have **: \<open>take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\<close> .
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   820
  have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (take_bit (Suc n) k + take_bit (Suc n) (2 ^ n))\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   821
    by (simp only: take_bit_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   822
  also have \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   823
    by (simp add: take_bit_Suc_from_most)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   824
  finally have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (2 ^ (n + of_bool (bit k n)) + take_bit n k)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   825
    by (simp add: ac_simps)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   826
  also have \<open>2 ^ (n + of_bool (bit k n)) + take_bit n k = 2 ^ (n + of_bool (bit k n)) OR take_bit n k\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   827
    by (rule disjunctive_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   828
      (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   829
  finally show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   830
    using * ** by (simp add: signed_take_bit_def take_bit_Suc min_def ac_simps)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   831
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   832
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   833
lemma signed_take_bit_of_0 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   834
  \<open>signed_take_bit n 0 = 0\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   835
  by (simp add: signed_take_bit_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   836
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   837
lemma signed_take_bit_of_minus_1 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   838
  \<open>signed_take_bit n (- 1) = - 1\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   839
  by (simp add: signed_take_bit_def take_bit_minus_one_eq_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   840
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   841
lemma signed_take_bit_eq_iff_take_bit_eq:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   842
  \<open>signed_take_bit n k = signed_take_bit n l \<longleftrightarrow> take_bit (Suc n) k = take_bit (Suc n) l\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   843
proof (cases \<open>bit k n \<longleftrightarrow> bit l n\<close>)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   844
  case True
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   845
  moreover have \<open>take_bit n k OR NOT (mask n) = take_bit n k - 2 ^ n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   846
    for k :: int
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   847
    by (auto simp add: disjunctive_add [symmetric] bit_not_iff bit_mask_iff bit_take_bit_iff minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   848
  ultimately show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   849
    by (simp add: signed_take_bit_def take_bit_Suc_from_most)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   850
next
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   851
  case False
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   852
  then have \<open>signed_take_bit n k \<noteq> signed_take_bit n l\<close> and \<open>take_bit (Suc n) k \<noteq> take_bit (Suc n) l\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   853
    by (auto simp add: bit_eq_iff bit_take_bit_iff bit_signed_take_bit_iff min_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   854
  then show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   855
    by simp
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   856
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   857
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   858
lemma signed_take_bit_signed_take_bit [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   859
  \<open>signed_take_bit m (signed_take_bit n k) = signed_take_bit (min m n) k\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   860
  by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_or_iff bit_not_iff bit_mask_iff bit_take_bit_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   861
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   862
lemma take_bit_signed_take_bit:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   863
  \<open>take_bit m (signed_take_bit n k) = take_bit m k\<close> if \<open>m \<le> n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   864
  using that by (auto simp add: bit_eq_iff bit_take_bit_iff bit_signed_take_bit_iff min_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   865
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   866
lemma take_bit_Suc_signed_take_bit [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   867
  \<open>take_bit (Suc n) (signed_take_bit n a) = take_bit (Suc n) a\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   868
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   869
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   870
lemma signed_take_bit_take_bit:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   871
  \<open>signed_take_bit m (take_bit n k) = (if n \<le> m then take_bit n else signed_take_bit m) k\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   872
  by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   873
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   874
lemma signed_take_bit_nonnegative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   875
  \<open>0 \<le> signed_take_bit n k \<longleftrightarrow> \<not> bit k n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   876
  by (simp add: signed_take_bit_def not_less mask_nonnegative)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   877
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   878
lemma signed_take_bit_negative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   879
  \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   880
  by (simp add: signed_take_bit_def not_less mask_nonnegative)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   881
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   882
lemma signed_take_bit_greater_eq:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   883
  \<open>k + 2 ^ Suc n \<le> signed_take_bit n k\<close> if \<open>k < - (2 ^ n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   884
  using that take_bit_greater_eq [of \<open>k + 2 ^ n\<close> \<open>Suc n\<close>]
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   885
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   886
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   887
lemma signed_take_bit_less_eq:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   888
  \<open>signed_take_bit n k \<le> k - 2 ^ Suc n\<close> if \<open>k \<ge> 2 ^ n\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   889
  using that take_bit_less_eq [of \<open>Suc n\<close> \<open>k + 2 ^ n\<close>]
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   890
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   891
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   892
lemma signed_take_bit_Suc_1 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   893
  \<open>signed_take_bit (Suc n) 1 = 1\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   894
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   895
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   896
lemma signed_take_bit_Suc_bit0 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   897
  \<open>signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * 2\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   898
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   899
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   900
lemma signed_take_bit_Suc_bit1 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   901
  \<open>signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + 1\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   902
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   903
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   904
lemma signed_take_bit_Suc_minus_bit0 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   905
  \<open>signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * 2\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   906
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   907
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   908
lemma signed_take_bit_Suc_minus_bit1 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   909
  \<open>signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + 1\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   910
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   911
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   912
lemma signed_take_bit_numeral_bit0 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   913
  \<open>signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   914
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   915
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   916
lemma signed_take_bit_numeral_bit1 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   917
  \<open>signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + 1\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   918
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   919
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   920
lemma signed_take_bit_numeral_minus_bit0 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   921
  \<open>signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * 2\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   922
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   923
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   924
lemma signed_take_bit_numeral_minus_bit1 [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   925
  \<open>signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + 1\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   926
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   927
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   928
lemma signed_take_bit_code [code]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   929
  \<open>signed_take_bit n k =
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   930
  (let l = k AND mask (Suc n)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   931
   in if bit l n then l - (push_bit n 2) else l)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   932
proof -
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   933
  have *: \<open>(k AND mask (Suc n)) - 2 * 2 ^ n = k AND mask (Suc n) OR NOT (mask (Suc n))\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   934
    apply (subst disjunctive_add [symmetric])
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   935
    apply (simp_all add: bit_and_iff bit_mask_iff bit_not_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   936
    apply (simp flip: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   937
    done
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   938
  show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   939
    by (rule bit_eqI)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   940
     (auto simp add: Let_def bit_and_iff bit_signed_take_bit_iff push_bit_eq_mult min_def not_le bit_mask_iff bit_exp_iff less_Suc_eq * bit_or_iff bit_not_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   941
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   942
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   943
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
   944
subsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   945
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   946
instantiation nat :: semiring_bit_operations
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   947
begin
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   948
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   949
definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   950
  where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   951
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   952
definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   953
  where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   954
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   955
definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   956
  where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   957
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   958
instance proof
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   959
  fix m n q :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   960
  show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   961
    by (auto simp add: and_nat_def bit_and_iff less_le bit_eq_iff)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   962
  show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   963
    by (auto simp add: or_nat_def bit_or_iff less_le bit_eq_iff)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   964
  show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   965
    by (auto simp add: xor_nat_def bit_xor_iff less_le bit_eq_iff)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   966
qed
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   967
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   968
end
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   969
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   970
lemma and_nat_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   971
  \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   972
  by (simp add: and_nat_def and_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   973
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   974
lemma or_nat_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   975
  \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   976
  by (simp add: or_nat_def or_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   977
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   978
lemma xor_nat_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   979
  \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   980
  by (simp add: xor_nat_def xor_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   981
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   982
lemma Suc_0_and_eq [simp]:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   983
  \<open>Suc 0 AND n = n mod 2\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   984
  using one_and_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   985
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   986
lemma and_Suc_0_eq [simp]:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   987
  \<open>n AND Suc 0 = n mod 2\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   988
  using and_one_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   989
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   990
lemma Suc_0_or_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   991
  \<open>Suc 0 OR n = n + of_bool (even n)\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   992
  using one_or_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   993
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   994
lemma or_Suc_0_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   995
  \<open>n OR Suc 0 = n + of_bool (even n)\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   996
  using or_one_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   997
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
   998
lemma Suc_0_xor_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   999
  \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1000
  using one_xor_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1001
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1002
lemma xor_Suc_0_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1003
  \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1004
  using xor_one_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1005
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1006
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1007
subsection \<open>Instances for \<^typ>\<open>integer\<close> and \<^typ>\<open>natural\<close>\<close>
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1008
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1009
unbundle integer.lifting natural.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1010
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1011
instantiation integer :: ring_bit_operations
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1012
begin
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1013
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1014
lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1015
  is not .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1016
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1017
lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1018
  is \<open>and\<close> .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1019
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1020
lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1021
  is or .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1022
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1023
lift_definition xor_integer ::  \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1024
  is xor .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1025
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1026
instance proof
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1027
  fix k l :: \<open>integer\<close> and n :: nat
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1028
  show \<open>- k = NOT (k - 1)\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1029
    by transfer (simp add: minus_eq_not_minus_1)
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1030
  show \<open>bit (NOT k) n \<longleftrightarrow> (2 :: integer) ^ n \<noteq> 0 \<and> \<not> bit k n\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1031
    by transfer (fact bit_not_iff)
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1032
  show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1033
    by transfer (fact bit_and_iff)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1034
  show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1035
    by transfer (fact bit_or_iff)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1036
  show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1037
    by transfer (fact bit_xor_iff)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1038
qed
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1039
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1040
end
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1041
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1042
instantiation natural :: semiring_bit_operations
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1043
begin
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1044
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1045
lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1046
  is \<open>and\<close> .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1047
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1048
lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1049
  is or .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1050
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1051
lift_definition xor_natural ::  \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1052
  is xor .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1053
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1054
instance proof
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1055
  fix m n :: \<open>natural\<close> and q :: nat
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1056
  show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1057
    by transfer (fact bit_and_iff)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1058
  show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1059
    by transfer (fact bit_or_iff)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1060
  show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1061
    by transfer (fact bit_xor_iff)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1062
qed
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1063
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1064
end
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1065
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1066
lifting_update integer.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1067
lifting_forget integer.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1068
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1069
lifting_update natural.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1070
lifting_forget natural.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1071
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1072
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1073
subsection \<open>Key ideas of bit operations\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1074
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1075
text \<open>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1076
  When formalizing bit operations, it is tempting to represent
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1077
  bit values as explicit lists over a binary type. This however
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1078
  is a bad idea, mainly due to the inherent ambiguities in
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1079
  representation concerning repeating leading bits.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1080
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1081
  Hence this approach avoids such explicit lists altogether
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1082
  following an algebraic path:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1083
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1084
  \<^item> Bit values are represented by numeric types: idealized
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1085
    unbounded bit values can be represented by type \<^typ>\<open>int\<close>,
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1086
    bounded bit values by quotient types over \<^typ>\<open>int\<close>.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1087
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1088
  \<^item> (A special case are idealized unbounded bit values ending
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1089
    in @{term [source] 0} which can be represented by type \<^typ>\<open>nat\<close> but
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1090
    only support a restricted set of operations).
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1091
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1092
  \<^item> From this idea follows that
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1093
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1094
      \<^item> multiplication by \<^term>\<open>2 :: int\<close> is a bit shift to the left and
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1095
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1096
      \<^item> division by \<^term>\<open>2 :: int\<close> is a bit shift to the right.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1097
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1098
  \<^item> Concerning bounded bit values, iterated shifts to the left
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1099
    may result in eliminating all bits by shifting them all
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1100
    beyond the boundary.  The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1101
    represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1102
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
  1103
  \<^item> The projection on a single bit is then @{thm bit_iff_odd [where ?'a = int, no_vars]}.
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1104
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1105
  \<^item> This leads to the most fundamental properties of bit values:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1106
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1107
      \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1108
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1109
      \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1110
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1111
  \<^item> Typical operations are characterized as follows:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1112
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1113
      \<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1114
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1115
      \<^item> Bit mask upto bit \<^term>\<open>n\<close>: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]}
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1116
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1117
      \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1118
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1119
      \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1120
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1121
      \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1122
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1123
      \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1124
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1125
      \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1126
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1127
      \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1128
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1129
      \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1130
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1131
      \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1132
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1133
      \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1134
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1135
      \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1136
\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1137
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1138
end