src/HOL/Library/Bit_Operations.thy
author haftmann
Wed, 23 Jun 2021 17:43:31 +0000
changeset 73868 465846b611d5
parent 73816 0510c7a4256a
child 73869 7181130f5872
permissions -rw-r--r--
some word streamlining
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
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
     8
    Main
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     9
    "HOL-Library.Boolean_Algebra"
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
73816
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
    12
lemma bit_numeral_int_iff [bit_simps]: \<comment> \<open>TODO: move\<close>
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
    13
  \<open>bit (numeral m :: int) n \<longleftrightarrow> bit (numeral m :: nat) n\<close>
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
    14
  using bit_of_nat_iff_bit [of \<open>numeral m\<close> n] by simp
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
    15
0510c7a4256a moved more legacy to AFP
haftmann
parents: 73789
diff changeset
    16
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
    17
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
    18
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    19
class semiring_bit_operations = semiring_bit_shifts +
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    20
  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
    21
    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
    22
    and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>XOR\<close> 59)
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
    23
    and mask :: \<open>nat \<Rightarrow> 'a\<close>
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    24
    and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    25
    and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    26
    and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    27
  assumes bit_and_iff [bit_simps]: \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    28
    and bit_or_iff [bit_simps]: \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    29
    and bit_xor_iff [bit_simps]: \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
    30
    and mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close>
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    31
    and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    32
    and bit_unset_bit_iff [bit_simps]: \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
    33
    and flip_bit_eq_xor: \<open>flip_bit n a = a XOR push_bit n 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
    34
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    35
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    36
text \<open>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    37
  We want the bitwise operations to bind slightly weaker
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71042
diff changeset
    38
  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
    39
  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
    40
  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
    41
  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
    42
\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    43
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    44
sublocale "and": semilattice \<open>(AND)\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    45
  by standard (auto simp add: bit_eq_iff bit_and_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    46
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    47
sublocale or: semilattice_neutr \<open>(OR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    48
  by standard (auto simp add: bit_eq_iff bit_or_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    49
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    50
sublocale xor: comm_monoid \<open>(XOR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    51
  by standard (auto simp add: bit_eq_iff bit_xor_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    52
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    53
lemma even_and_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    54
  \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    55
  using bit_and_iff [of a b 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    56
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    57
lemma even_or_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    58
  \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    59
  using bit_or_iff [of a b 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    60
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    61
lemma even_xor_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    62
  \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    63
  using bit_xor_iff [of a b 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
    64
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    65
lemma zero_and_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    66
  "0 AND a = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    67
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    68
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    69
lemma and_zero_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    70
  "a AND 0 = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    71
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    72
71921
a238074c5a9d avoid overaggressive default simp rules
haftmann
parents: 71823
diff changeset
    73
lemma one_and_eq:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    74
  "1 AND a = a mod 2"
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    75
  by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    76
71921
a238074c5a9d avoid overaggressive default simp rules
haftmann
parents: 71823
diff changeset
    77
lemma and_one_eq:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    78
  "a AND 1 = a mod 2"
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    79
  using one_and_eq [of a] by (simp add: ac_simps)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    80
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    81
lemma one_or_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    82
  "1 OR a = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    83
  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
    84
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    85
lemma or_one_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    86
  "a OR 1 = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    87
  using one_or_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    88
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    89
lemma one_xor_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    90
  "1 XOR a = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    91
  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
    92
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
    93
lemma xor_one_eq:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    94
  "a XOR 1 = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    95
  using one_xor_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    96
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
    97
lemma take_bit_and [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
    98
  \<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
    99
  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
   100
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   101
lemma take_bit_or [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   102
  \<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
   103
  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
   104
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   105
lemma take_bit_xor [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   106
  \<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
   107
  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
   108
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   109
lemma push_bit_and [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   110
  \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   111
  by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   112
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   113
lemma push_bit_or [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   114
  \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   115
  by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   116
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   117
lemma push_bit_xor [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   118
  \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   119
  by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   120
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   121
lemma drop_bit_and [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   122
  \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   123
  by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   124
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   125
lemma drop_bit_or [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   126
  \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   127
  by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   128
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   129
lemma drop_bit_xor [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   130
  \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   131
  by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   132
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   133
lemma bit_mask_iff [bit_simps]:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   134
  \<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
   135
  by (simp add: mask_eq_exp_minus_1 bit_mask_iff)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   136
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   137
lemma even_mask_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   138
  \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   139
  using bit_mask_iff [of n 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   140
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   141
lemma mask_0 [simp]:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   142
  \<open>mask 0 = 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   143
  by (simp add: mask_eq_exp_minus_1)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   144
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   145
lemma mask_Suc_0 [simp]:
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   146
  \<open>mask (Suc 0) = 1\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   147
  by (simp add: mask_eq_exp_minus_1 add_implies_diff sym)
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   148
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   149
lemma mask_Suc_exp:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   150
  \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   151
  by (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   152
    (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
   153
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   154
lemma mask_Suc_double:
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   155
  \<open>mask (Suc n) = 1 OR 2 * mask n\<close>
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   156
proof (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   157
  fix q
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   158
  assume \<open>2 ^ q \<noteq> 0\<close>
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   159
  show \<open>bit (mask (Suc n)) q \<longleftrightarrow> bit (1 OR 2 * mask n) q\<close>
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   160
    by (cases q)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   161
      (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
   162
qed
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   163
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   164
lemma mask_numeral:
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   165
  \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   166
  by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps)
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   167
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
   168
lemma take_bit_mask [simp]:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
   169
  \<open>take_bit m (mask n) = mask (min m n)\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
   170
  by (rule bit_eqI) (simp add: bit_simps)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
   171
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
   172
lemma take_bit_eq_mask:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   173
  \<open>take_bit n a = a AND mask n\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   174
  by (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   175
    (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
   176
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   177
lemma or_eq_0_iff:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   178
  \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72611
diff changeset
   179
  by (auto simp add: bit_eq_iff bit_or_iff)
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   180
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   181
lemma disjunctive_add:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   182
  \<open>a + b = a OR b\<close> if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   183
  by (rule bit_eqI) (use that in \<open>simp add: bit_disjunctive_add_iff bit_or_iff\<close>)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   184
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   185
lemma bit_iff_and_drop_bit_eq_1:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   186
  \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   187
  by (simp add: bit_iff_odd_drop_bit and_one_eq odd_iff_mod_2_eq_one)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   188
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   189
lemma bit_iff_and_push_bit_not_eq_0:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   190
  \<open>bit a n \<longleftrightarrow> a AND push_bit n 1 \<noteq> 0\<close>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   191
  apply (cases \<open>2 ^ n = 0\<close>)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   192
  apply (simp_all add: push_bit_of_1 bit_eq_iff bit_and_iff bit_push_bit_iff exp_eq_0_imp_not_bit)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   193
  apply (simp_all add: bit_exp_iff)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   194
  done
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
   195
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   196
lemmas set_bit_def = set_bit_eq_or
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   197
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   198
lemma bit_set_bit_iff [bit_simps]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   199
  \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> 2 ^ n \<noteq> 0)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   200
  by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   201
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   202
lemma even_set_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   203
  \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   204
  using bit_set_bit_iff [of m a 0] by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   205
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   206
lemma even_unset_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   207
  \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   208
  using bit_unset_bit_iff [of m a 0] by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   209
73789
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   210
lemma and_exp_eq_0_iff_not_bit:
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   211
  \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   212
proof
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   213
  assume ?Q
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   214
  then show ?P
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   215
    by (auto intro: bit_eqI simp add: bit_simps)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   216
next
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   217
  assume ?P
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   218
  show ?Q
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   219
  proof (rule notI)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   220
    assume \<open>bit a n\<close>
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   221
    then have \<open>a AND 2 ^ n = 2 ^ n\<close>
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   222
      by (auto intro: bit_eqI simp add: bit_simps)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   223
    with \<open>?P\<close> show False
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   224
      using \<open>bit a n\<close> exp_eq_0_imp_not_bit by auto
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   225
  qed
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   226
qed
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
   227
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   228
lemmas flip_bit_def = flip_bit_eq_xor
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   229
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   230
lemma bit_flip_bit_iff [bit_simps]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   231
  \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> 2 ^ n \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   232
  by (auto simp add: flip_bit_def push_bit_of_1 bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   233
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   234
lemma even_flip_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   235
  \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   236
  using bit_flip_bit_iff [of m a 0] by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   237
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   238
lemma set_bit_0 [simp]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   239
  \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   240
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   241
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   242
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   243
  then show \<open>bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   244
    by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   245
      (cases m, simp_all add: bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   246
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   247
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   248
lemma set_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   249
  \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   250
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   251
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   252
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   253
  show \<open>bit (set_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * set_bit n (a div 2)) m\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   254
  proof (cases m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   255
    case 0
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   256
    then show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   257
      by (simp add: even_set_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   258
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   259
    case (Suc m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   260
    with * have \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   261
      using mult_2 by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   262
    show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   263
      by (cases a rule: parity_cases)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   264
        (simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *,
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   265
        simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   266
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   267
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   268
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   269
lemma unset_bit_0 [simp]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   270
  \<open>unset_bit 0 a = 2 * (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   271
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   272
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   273
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   274
  then show \<open>bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   275
    by (simp add: bit_unset_bit_iff bit_double_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   276
      (cases m, simp_all add: bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   277
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   278
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   279
lemma unset_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   280
  \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   281
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   282
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   283
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   284
  then show \<open>bit (unset_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * unset_bit n (a div 2)) m\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   285
  proof (cases m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   286
    case 0
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   287
    then show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   288
      by (simp add: even_unset_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   289
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   290
    case (Suc m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   291
    show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   292
      by (cases a rule: parity_cases)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   293
        (simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *,
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   294
         simp_all add: Suc bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   295
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   296
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   297
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   298
lemma flip_bit_0 [simp]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   299
  \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   300
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   301
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   302
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   303
  then show \<open>bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   304
    by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   305
      (cases m, simp_all add: bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   306
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   307
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   308
lemma flip_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   309
  \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   310
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   311
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   312
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   313
  show \<open>bit (flip_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * flip_bit n (a div 2)) m\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   314
  proof (cases m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   315
    case 0
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   316
    then show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   317
      by (simp add: even_flip_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   318
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   319
    case (Suc m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   320
    with * have \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   321
      using mult_2 by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   322
    show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   323
      by (cases a rule: parity_cases)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   324
        (simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff,
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   325
        simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   326
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   327
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   328
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   329
lemma flip_bit_eq_if:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   330
  \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   331
  by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   332
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   333
lemma take_bit_set_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   334
  \<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>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   335
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   336
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   337
lemma take_bit_unset_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   338
  \<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>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   339
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   340
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   341
lemma take_bit_flip_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   342
  \<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>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   343
  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   344
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   345
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   346
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   347
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   348
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
   349
  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   350
  assumes bit_not_iff [bit_simps]: \<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
   351
  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
   352
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   353
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   354
text \<open>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   355
  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
   356
  definitional class operation.  Note that \<^const>\<open>not\<close> has no
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   357
  sensible definition for unlimited but only positive bit strings
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   358
  (type \<^typ>\<open>nat\<close>).
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   359
\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   360
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   361
lemma bits_minus_1_mod_2_eq [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   362
  \<open>(- 1) mod 2 = 1\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   363
  by (simp add: mod_2_eq_odd)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   364
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   365
lemma not_eq_complement:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   366
  \<open>NOT a = - a - 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   367
  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
   368
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   369
lemma minus_eq_not_plus_1:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   370
  \<open>- a = NOT a + 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   371
  using not_eq_complement [of a] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   372
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   373
lemma bit_minus_iff [bit_simps]:
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   374
  \<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
   375
  by (simp add: minus_eq_not_minus_1 bit_not_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   376
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   377
lemma even_not_iff [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   378
  "even (NOT a) \<longleftrightarrow> odd a"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   379
  using bit_not_iff [of a 0] by auto
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   380
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   381
lemma bit_not_exp_iff [bit_simps]:
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   382
  \<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
   383
  by (auto simp add: bit_not_iff bit_exp_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   384
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   385
lemma bit_minus_1_iff [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   386
  \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   387
  by (simp add: bit_minus_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   388
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   389
lemma bit_minus_exp_iff [bit_simps]:
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   390
  \<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close>
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   391
  by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   392
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   393
lemma bit_minus_2_iff [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   394
  \<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
   395
  by (simp add: bit_minus_iff bit_1_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   396
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   397
lemma not_one [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   398
  "NOT 1 = - 2"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   399
  by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   400
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   401
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   402
  by standard (rule bit_eqI, simp add: bit_and_iff)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   403
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   404
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
   405
  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
   406
proof -
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   407
  interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   408
    by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   409
  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
   410
    by standard
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   411
  show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   412
    by (rule ext, rule ext, rule bit_eqI)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   413
      (auto simp add: bit.xor_def bit_and_iff bit_or_iff bit_xor_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
   414
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   415
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   416
lemma and_eq_not_not_or:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   417
  \<open>a AND b = NOT (NOT a OR NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   418
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   419
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   420
lemma or_eq_not_not_and:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   421
  \<open>a OR b = NOT (NOT a AND NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   422
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   423
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   424
lemma not_add_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   425
  \<open>NOT (a + b) = NOT a - b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   426
  by (simp add: not_eq_complement algebra_simps)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   427
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   428
lemma not_diff_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   429
  \<open>NOT (a - b) = NOT a + b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   430
  using not_add_distrib [of a \<open>- b\<close>] by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   431
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   432
lemma (in ring_bit_operations) and_eq_minus_1_iff:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   433
  \<open>a AND b = - 1 \<longleftrightarrow> a = - 1 \<and> b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   434
proof
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   435
  assume \<open>a = - 1 \<and> b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   436
  then show \<open>a AND b = - 1\<close>
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72611
diff changeset
   437
    by simp
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   438
next
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   439
  assume \<open>a AND b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   440
  have *: \<open>bit a n\<close> \<open>bit b n\<close> if \<open>2 ^ n \<noteq> 0\<close> for n
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   441
  proof -
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   442
    from \<open>a AND b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   443
    have \<open>bit (a AND b) n = bit (- 1) n\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   444
      by (simp add: bit_eq_iff)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   445
    then show \<open>bit a n\<close> \<open>bit b n\<close>
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72611
diff changeset
   446
      using that by (simp_all add: bit_and_iff)
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   447
  qed
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   448
  have \<open>a = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   449
    by (rule bit_eqI) (simp add: *)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   450
  moreover have \<open>b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   451
    by (rule bit_eqI) (simp add: *)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   452
  ultimately show \<open>a = - 1 \<and> b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   453
    by simp
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   454
qed
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
   455
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   456
lemma disjunctive_diff:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   457
  \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   458
proof -
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   459
  have \<open>NOT a + b = NOT a OR b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   460
    by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   461
  then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   462
    by simp
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   463
  then show ?thesis
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   464
    by (simp add: not_add_distrib)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   465
qed
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   466
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   467
lemma push_bit_minus:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   468
  \<open>push_bit n (- a) = - push_bit n a\<close>
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   469
  by (simp add: push_bit_eq_mult)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   470
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   471
lemma take_bit_not_take_bit:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   472
  \<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
   473
  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
   474
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   475
lemma take_bit_not_iff:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   476
  "take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b"
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   477
  apply (simp add: bit_eq_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   478
  apply (simp add: bit_not_iff bit_take_bit_iff bit_exp_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
   479
  apply (use exp_eq_0_imp_not_bit in blast)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   480
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   481
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   482
lemma take_bit_not_eq_mask_diff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   483
  \<open>take_bit n (NOT a) = mask n - take_bit n a\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   484
proof -
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   485
  have \<open>take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   486
    by (simp add: take_bit_not_take_bit)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   487
  also have \<open>\<dots> = mask n AND NOT (take_bit n a)\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   488
    by (simp add: take_bit_eq_mask ac_simps)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   489
  also have \<open>\<dots> = mask n - take_bit n a\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   490
    by (subst disjunctive_diff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   491
      (auto simp add: bit_take_bit_iff bit_mask_iff exp_eq_0_imp_not_bit)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   492
  finally show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   493
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   494
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
   495
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
   496
lemma mask_eq_take_bit_minus_one:
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
   497
  \<open>mask n = take_bit n (- 1)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
   498
  by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
   499
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   500
lemma take_bit_minus_one_eq_mask:
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   501
  \<open>take_bit n (- 1) = mask n\<close>
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
   502
  by (simp add: mask_eq_take_bit_minus_one)
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   503
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   504
lemma minus_exp_eq_not_mask:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   505
  \<open>- (2 ^ n) = NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   506
  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
   507
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   508
lemma push_bit_minus_one_eq_not_mask:
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
   509
  \<open>push_bit n (- 1) = NOT (mask n)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   510
  by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   511
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   512
lemma take_bit_not_mask_eq_0:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   513
  \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
   514
  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
   515
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   516
lemma unset_bit_eq_and_not:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   517
  \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   518
  by (rule bit_eqI) (auto simp add: bit_simps)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
   519
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   520
lemmas unset_bit_def = unset_bit_eq_and_not
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   521
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   522
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   523
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   524
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
   525
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
   526
72397
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   527
lemma int_bit_bound:
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   528
  fixes k :: int
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   529
  obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   530
    and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   531
proof -
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   532
  obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   533
  proof (cases \<open>k \<ge> 0\<close>)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   534
    case True
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   535
    moreover from power_gt_expt [of 2 \<open>nat k\<close>]
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   536
    have \<open>k < 2 ^ nat k\<close> by simp
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   537
    ultimately have *: \<open>k div 2 ^ nat k = 0\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   538
      by simp
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   539
    show thesis
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   540
    proof (rule that [of \<open>nat k\<close>])
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   541
      fix m
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   542
      assume \<open>nat k \<le> m\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   543
      then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   544
        by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   545
    qed
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   546
  next
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   547
    case False
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   548
    moreover from power_gt_expt [of 2 \<open>nat (- k)\<close>]
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   549
    have \<open>- k \<le> 2 ^ nat (- k)\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   550
      by simp
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   551
    ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   552
      by (subst div_pos_neg_trivial) simp_all
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   553
    then have *: \<open>k div 2 ^ nat (- k) = - 1\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   554
      by simp
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   555
    show thesis
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   556
    proof (rule that [of \<open>nat (- k)\<close>])
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   557
      fix m
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   558
      assume \<open>nat (- k) \<le> m\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   559
      then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   560
        by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   561
    qed
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   562
  qed
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   563
  show thesis
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   564
  proof (cases \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   565
    case True
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   566
    then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   567
      by blast
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   568
    with True that [of 0] show thesis
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   569
      by simp
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   570
  next
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   571
    case False
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   572
    then obtain r where **: \<open>bit k r \<noteq> bit k q\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   573
      by blast
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   574
    have \<open>r < q\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   575
      by (rule ccontr) (use * [of r] ** in simp)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   576
    define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   577
    moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   578
      using ** N_def \<open>r < q\<close> by auto
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   579
    moreover define n where \<open>n = Suc (Max N)\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   580
    ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   581
      apply auto
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   582
         apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   583
        apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   584
        apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   585
      apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   586
      done
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   587
    have \<open>bit k (Max N) \<noteq> bit k n\<close>
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   588
      by (metis (mono_tags, lifting) "*" Max_in N_def \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> \<open>finite N\<close> \<open>r \<in> N\<close> empty_iff le_cases mem_Collect_eq)
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   589
    show thesis apply (rule that [of n])
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   590
      using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   591
      using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   592
  qed
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   593
qed
48013583e8e6 factored out bit comprehension
haftmann
parents: 72281
diff changeset
   594
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   595
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
   596
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   597
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   598
definition not_int :: \<open>int \<Rightarrow> int\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   599
  where \<open>not_int k = - k - 1\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   600
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   601
lemma not_int_rec:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   602
  "NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   603
  by (auto simp add: not_int_def elim: oddE)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   604
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   605
lemma even_not_iff_int:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   606
  \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   607
  by (simp add: not_int_def)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   608
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   609
lemma not_int_div_2:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   610
  \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   611
  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
   612
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
   613
lemma bit_not_int_iff [bit_simps]:
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   614
  \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   615
  for k :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   616
  by (simp add: bit_not_int_iff' not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   617
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   618
function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   619
  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
   620
    then - of_bool (odd k \<and> odd l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   621
    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
   622
  by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   623
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   624
termination
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   625
  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
   626
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   627
declare and_int.simps [simp del]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   628
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   629
lemma and_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   630
  \<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
   631
    for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   632
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
   633
  case True
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   634
  then show ?thesis
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   635
    by auto (simp_all add: and_int.simps)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   636
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   637
  case False
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   638
  then show ?thesis
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   639
    by (auto simp add: ac_simps and_int.simps [of k l])
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   640
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   641
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   642
lemma bit_and_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   643
  \<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
   644
proof (induction n arbitrary: k l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   645
  case 0
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   646
  then show ?case
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   647
    by (simp add: and_int_rec [of k l])
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   648
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   649
  case (Suc n)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   650
  then show ?case
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   651
    by (simp add: and_int_rec [of k l] bit_Suc)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   652
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   653
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   654
lemma even_and_iff_int:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   655
  \<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
   656
  using bit_and_int_iff [of k l 0] by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   657
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   658
definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   659
  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
   660
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   661
lemma or_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   662
  \<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
   663
  for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   664
  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
   665
  by (simp add: or_int_def even_not_iff_int not_int_div_2)
73535
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   666
    (simp_all add: not_int_def)
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   667
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   668
lemma bit_or_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   669
  \<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
   670
  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
   671
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   672
definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   673
  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
   674
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   675
lemma xor_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   676
  \<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
   677
  for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   678
  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
   679
    (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
   680
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   681
lemma bit_xor_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   682
  \<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
   683
  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
   684
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   685
definition mask_int :: \<open>nat \<Rightarrow> int\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   686
  where \<open>mask n = (2 :: int) ^ n - 1\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
   687
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   688
definition set_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   689
  where \<open>set_bit n k = k OR push_bit n 1\<close> for k :: int
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   690
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   691
definition unset_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   692
  where \<open>unset_bit n k = k AND NOT (push_bit n 1)\<close> for k :: int
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   693
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   694
definition flip_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   695
  where \<open>flip_bit n k = k XOR push_bit n 1\<close> for k :: int
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   696
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   697
instance proof
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   698
  fix k l :: int and m n :: nat
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   699
  show \<open>- k = NOT (k - 1)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   700
    by (simp add: not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   701
  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
   702
    by (fact bit_and_int_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   703
  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
   704
    by (fact bit_or_int_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   705
  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
   706
    by (fact bit_xor_int_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   707
  show \<open>bit (unset_bit m k) n \<longleftrightarrow> bit k n \<and> m \<noteq> n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   708
  proof -
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   709
    have \<open>unset_bit m k = k AND NOT (push_bit m 1)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   710
      by (simp add: unset_bit_int_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   711
    also have \<open>NOT (push_bit m 1 :: int) = - (push_bit m 1 + 1)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   712
      by (simp add: not_int_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   713
    finally show ?thesis by (simp only: bit_simps bit_and_int_iff) (auto simp add: bit_simps)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   714
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
   715
qed (simp_all add: bit_not_int_iff mask_int_def set_bit_int_def flip_bit_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
   716
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   717
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   718
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   719
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
   720
lemma mask_half_int:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
   721
  \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
   722
  by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
   723
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   724
lemma mask_nonnegative_int [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   725
  \<open>mask n \<ge> (0::int)\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   726
  by (simp add: mask_eq_exp_minus_1)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   727
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   728
lemma not_mask_negative_int [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   729
  \<open>\<not> mask n < (0::int)\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   730
  by (simp add: not_less)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
   731
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   732
lemma not_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   733
  \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   734
  by (simp add: not_int_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   735
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   736
lemma not_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   737
  \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   738
  by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   739
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   740
lemma and_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   741
  \<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
   742
proof (induction k arbitrary: l rule: int_bit_induct)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   743
  case zero
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   744
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   745
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   746
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   747
  case minus
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   748
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   749
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   750
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   751
  case (even k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   752
  then show ?case
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   753
    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
   754
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   755
  case (odd k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   756
  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
   757
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   758
  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
   759
    by simp
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
   760
  with and_int_rec [of \<open>1 + k * 2\<close> l]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   761
  show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   762
    by auto
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   763
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   764
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   765
lemma and_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   766
  \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   767
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   768
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   769
lemma and_less_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   770
  \<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
   771
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   772
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   773
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   774
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   775
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   776
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   777
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   778
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   779
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   780
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   781
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   782
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   783
    by (simp add: and_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   784
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   785
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   786
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   787
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   788
    by (simp add: and_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   789
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   790
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   791
lemma or_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   792
  \<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
   793
  by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   794
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   795
lemma or_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   796
  \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   797
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   798
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   799
lemma or_greater_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   800
  \<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
   801
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   802
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   803
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   804
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   805
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   806
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   807
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   808
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   809
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   810
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   811
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   812
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   813
    by (simp add: or_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   814
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   815
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   816
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   817
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   818
    by (simp add: or_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   819
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   820
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   821
lemma xor_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   822
  \<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
   823
  by (simp only: bit.xor_def or_nonnegative_int_iff) auto
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   824
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   825
lemma xor_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   826
  \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   827
  by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   828
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   829
lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   830
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   831
  assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   832
  shows "x OR y < 2 ^ n"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   833
using assms proof (induction x arbitrary: y n rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   834
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   835
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   836
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   837
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   838
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   839
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   840
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   841
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   842
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   843
  from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   844
  show ?case 
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   845
    by (cases n) (auto simp add: or_int_rec [of \<open>_ * 2\<close>] elim: oddE)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   846
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   847
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   848
  from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   849
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   850
    by (cases n) (auto simp add: or_int_rec [of \<open>1 + _ * 2\<close>], linarith)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   851
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   852
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   853
lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   854
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   855
  assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   856
  shows "x XOR y < 2 ^ n"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   857
using assms proof (induction x arbitrary: y n rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   858
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   859
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   860
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   861
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   862
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   863
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   864
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   865
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   866
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   867
  from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   868
  show ?case 
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   869
    by (cases n) (auto simp add: xor_int_rec [of \<open>_ * 2\<close>] elim: oddE)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   870
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   871
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   872
  from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   873
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   874
    by (cases n) (auto simp add: xor_int_rec [of \<open>1 + _ * 2\<close>])
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   875
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   876
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   877
lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   878
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   879
  assumes "0 \<le> x"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   880
  shows "0 \<le> x AND y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   881
  using assms by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   882
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   883
lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   884
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   885
  assumes "0 \<le> x" "0 \<le> y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   886
  shows "0 \<le> x OR y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   887
  using assms by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   888
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   889
lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   890
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   891
  assumes "0 \<le> x" "0 \<le> y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   892
  shows "0 \<le> x XOR y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   893
  using assms by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   894
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   895
lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   896
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   897
  assumes "0 \<le> x"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   898
  shows "x AND y \<le> x"
73535
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   899
using assms proof (induction x arbitrary: y rule: int_bit_induct)
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   900
  case (odd k)
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   901
  then have \<open>k AND y div 2 \<le> k\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   902
    by simp
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   903
  then show ?case 
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   904
    by (simp add: and_int_rec [of \<open>1 + _ * 2\<close>])
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
   905
qed (simp_all add: and_int_rec [of \<open>_ * 2\<close>])
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   906
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   907
lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   908
lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   909
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   910
lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   911
  fixes x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   912
  assumes "0 \<le> y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   913
  shows "x AND y \<le> y"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   914
  using assms AND_upper1 [of y x] by (simp add: ac_simps)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   915
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   916
lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   917
lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   918
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   919
lemma plus_and_or: \<open>(x AND y) + (x OR y) = x + y\<close> for x y :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   920
proof (induction x arbitrary: y rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   921
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   922
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   923
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   924
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   925
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   926
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   927
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   928
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   929
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   930
  from even.IH [of \<open>y div 2\<close>]
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   931
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   932
    by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   933
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   934
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   935
  from odd.IH [of \<open>y div 2\<close>]
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   936
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   937
    by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   938
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
   939
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   940
lemma set_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   941
  \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   942
  by (simp add: set_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   943
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   944
lemma set_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   945
  \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   946
  by (simp add: set_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   947
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   948
lemma unset_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   949
  \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   950
  by (simp add: unset_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   951
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   952
lemma unset_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   953
  \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   954
  by (simp add: unset_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   955
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   956
lemma flip_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   957
  \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   958
  by (simp add: flip_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   959
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   960
lemma flip_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   961
  \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   962
  by (simp add: flip_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
   963
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   964
lemma set_bit_greater_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   965
  \<open>set_bit n k \<ge> k\<close> for k :: int
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   966
  by (simp add: set_bit_def or_greater_eq)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   967
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   968
lemma unset_bit_less_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   969
  \<open>unset_bit n k \<le> k\<close> for k :: int
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   970
  by (simp add: unset_bit_def and_less_eq)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
   971
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   972
lemma set_bit_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   973
  \<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
   974
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   975
  fix m
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   976
  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
   977
  proof (cases \<open>m = n\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   978
    case True
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   979
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   980
      apply (simp add: bit_set_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   981
      apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   982
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   983
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   984
    case False
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   985
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   986
      apply (clarsimp simp add: bit_set_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   987
      apply (subst disjunctive_add)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   988
      apply (clarsimp simp add: bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   989
      apply (clarsimp simp add: bit_or_iff bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   990
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   991
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   992
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   993
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   994
lemma unset_bit_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   995
  \<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
   996
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   997
  fix m
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
   998
  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
   999
  proof (cases \<open>m = n\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1000
    case True
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1001
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1002
      apply (simp add: bit_unset_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1003
      apply (simp add: bit_iff_odd)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1004
      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
  1005
      apply (simp add: dvd_neg_div)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1006
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1007
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1008
    case False
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1009
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1010
      apply (clarsimp simp add: bit_unset_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1011
      apply (subst disjunctive_diff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1012
      apply (clarsimp simp add: bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1013
      apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1014
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1015
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1016
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1017
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1018
lemma take_bit_eq_mask_iff:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1019
  \<open>take_bit n k = mask n \<longleftrightarrow> take_bit n (k + 1) = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1020
  for k :: int
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1021
proof
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1022
  assume ?P
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1023
  then have \<open>take_bit n (take_bit n k + take_bit n 1) = 0\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1024
    by (simp add: mask_eq_exp_minus_1)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1025
  then show ?Q
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1026
    by (simp only: take_bit_add)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1027
next
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1028
  assume ?Q
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1029
  then have \<open>take_bit n (k + 1) - 1 = - 1\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1030
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1031
  then have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1032
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1033
  moreover have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n k\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1034
    by (simp add: take_bit_eq_mod mod_simps)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1035
  ultimately show ?P
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1036
    by (simp add: take_bit_minus_one_eq_mask)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1037
qed
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1038
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1039
lemma take_bit_eq_mask_iff_exp_dvd:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1040
  \<open>take_bit n k = mask n \<longleftrightarrow> 2 ^ n dvd k + 1\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1041
  for k :: int
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1042
  by (simp add: take_bit_eq_mask_iff flip: take_bit_eq_0_iff)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1043
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1044
context ring_bit_operations
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1045
begin
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1046
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1047
lemma even_of_int_iff:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1048
  \<open>even (of_int k) \<longleftrightarrow> even k\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1049
  by (induction k rule: int_bit_induct) simp_all
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1050
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1051
lemma bit_of_int_iff [bit_simps]:
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1052
  \<open>bit (of_int k) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit k n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1053
proof (cases \<open>(2::'a) ^ n = 0\<close>)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1054
  case True
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1055
  then show ?thesis
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1056
    by (simp add: exp_eq_0_imp_not_bit)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1057
next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1058
  case False
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1059
  then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1060
  proof (induction k arbitrary: n rule: int_bit_induct)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1061
    case zero
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1062
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1063
      by simp
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1064
  next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1065
    case minus
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1066
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1067
      by simp
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1068
  next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1069
    case (even k)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1070
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1071
      using bit_double_iff [of \<open>of_int k\<close> n] Parity.bit_double_iff [of k n]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1072
      by (cases n) (auto simp add: ac_simps dest: mult_not_zero)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1073
  next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1074
    case (odd k)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1075
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1076
      using bit_double_iff [of \<open>of_int k\<close> n]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1077
      by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Parity.bit_Suc dest: mult_not_zero)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1078
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1079
  with False show ?thesis
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1080
    by simp
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1081
qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1082
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1083
lemma push_bit_of_int:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1084
  \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1085
  by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1086
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1087
lemma of_int_push_bit:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1088
  \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1089
  by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1090
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1091
lemma take_bit_of_int:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1092
  \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1093
  by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1094
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1095
lemma of_int_take_bit:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1096
  \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1097
  by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1098
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1099
lemma of_int_not_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1100
  \<open>of_int (NOT k) = NOT (of_int k)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1101
  by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1102
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1103
lemma of_int_and_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1104
  \<open>of_int (k AND l) = of_int k AND of_int l\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1105
  by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1106
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1107
lemma of_int_or_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1108
  \<open>of_int (k OR l) = of_int k OR of_int l\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1109
  by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1110
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1111
lemma of_int_xor_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1112
  \<open>of_int (k XOR l) = of_int k XOR of_int l\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1113
  by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1114
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1115
lemma of_int_mask_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1116
  \<open>of_int (mask n) = mask n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1117
  by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1118
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1119
end
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1120
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1121
text \<open>FIXME: The rule sets below are very large (24 rules for each
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1122
  operator). Is there a simpler way to do this?\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1123
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1124
context
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1125
begin
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1126
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1127
private lemma eqI:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1128
  \<open>k = l\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1129
  if num: \<open>\<And>n. bit k (numeral n) \<longleftrightarrow> bit l (numeral n)\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1130
    and even: \<open>even k \<longleftrightarrow> even l\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1131
  for k l :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1132
proof (rule bit_eqI)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1133
  fix n
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1134
  show \<open>bit k n \<longleftrightarrow> bit l n\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1135
  proof (cases n)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1136
    case 0
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1137
    with even show ?thesis
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1138
      by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1139
  next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1140
    case (Suc n)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1141
    with num [of \<open>num_of_nat (Suc n)\<close>] show ?thesis
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1142
      by (simp only: numeral_num_of_nat)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1143
  qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1144
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1145
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1146
lemma int_and_numerals [simp]:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1147
  "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1148
  "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1149
  "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1150
  "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1151
  "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1152
  "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1153
  "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1154
  "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1155
  "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1156
  "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1157
  "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1158
  "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1159
  "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1160
  "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1161
  "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1162
  "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1163
  "(1::int) AND numeral (Num.Bit0 y) = 0"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1164
  "(1::int) AND numeral (Num.Bit1 y) = 1"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1165
  "(1::int) AND - numeral (Num.Bit0 y) = 0"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1166
  "(1::int) AND - numeral (Num.Bit1 y) = 1"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1167
  "numeral (Num.Bit0 x) AND (1::int) = 0"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1168
  "numeral (Num.Bit1 x) AND (1::int) = 1"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1169
  "- numeral (Num.Bit0 x) AND (1::int) = 0"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1170
  "- numeral (Num.Bit1 x) AND (1::int) = 1"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1171
  by (auto simp add: bit_and_iff bit_minus_iff even_and_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq intro: eqI)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1172
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1173
lemma int_or_numerals [simp]:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1174
  "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1175
  "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1176
  "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1177
  "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1178
  "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1179
  "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1180
  "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1181
  "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1182
  "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1183
  "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1184
  "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1185
  "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1186
  "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1187
  "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1188
  "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1189
  "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1190
  "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1191
  "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1192
  "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1193
  "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1194
  "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1195
  "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1196
  "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1197
  "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1198
  by (auto simp add: bit_or_iff bit_minus_iff even_or_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1199
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1200
lemma int_xor_numerals [simp]:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1201
  "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1202
  "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1203
  "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1204
  "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1205
  "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1206
  "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1207
  "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1208
  "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1209
  "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1210
  "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1211
  "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1212
  "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1213
  "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1214
  "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1215
  "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1216
  "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1217
  "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1218
  "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1219
  "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1220
  "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1221
  "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1222
  "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1223
  "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1224
  "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))"
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1225
  by (auto simp add: bit_xor_iff bit_minus_iff even_xor_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1226
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1227
end
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1228
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1229
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1230
subsection \<open>Bit concatenation\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1231
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1232
definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close>
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1233
  where \<open>concat_bit n k l = take_bit n k OR push_bit n l\<close>
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1234
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1235
lemma bit_concat_bit_iff [bit_simps]:
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1236
  \<open>bit (concat_bit m k l) n \<longleftrightarrow> n < m \<and> bit k n \<or> m \<le> n \<and> bit l (n - m)\<close>
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1237
  by (simp add: concat_bit_def bit_or_iff bit_and_iff bit_take_bit_iff bit_push_bit_iff ac_simps)
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1238
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1239
lemma concat_bit_eq:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1240
  \<open>concat_bit n k l = take_bit n k + push_bit n l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1241
  by (simp add: concat_bit_def take_bit_eq_mask
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1242
    bit_and_iff bit_mask_iff bit_push_bit_iff disjunctive_add)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1243
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1244
lemma concat_bit_0 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1245
  \<open>concat_bit 0 k l = l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1246
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1247
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1248
lemma concat_bit_Suc:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1249
  \<open>concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1250
  by (simp add: concat_bit_eq take_bit_Suc push_bit_double)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1251
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1252
lemma concat_bit_of_zero_1 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1253
  \<open>concat_bit n 0 l = push_bit n l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1254
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1255
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1256
lemma concat_bit_of_zero_2 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1257
  \<open>concat_bit n k 0 = take_bit n k\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1258
  by (simp add: concat_bit_def take_bit_eq_mask)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1259
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1260
lemma concat_bit_nonnegative_iff [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1261
  \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1262
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1263
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1264
lemma concat_bit_negative_iff [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1265
  \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1266
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1267
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1268
lemma concat_bit_assoc:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1269
  \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1270
  by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1271
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1272
lemma concat_bit_assoc_sym:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1273
  \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1274
  by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1275
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1276
lemma concat_bit_eq_iff:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1277
  \<open>concat_bit n k l = concat_bit n r s
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1278
    \<longleftrightarrow> take_bit n k = take_bit n r \<and> l = s\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1279
proof
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1280
  assume ?Q
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1281
  then show ?P
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1282
    by (simp add: concat_bit_def)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1283
next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1284
  assume ?P
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1285
  then have *: \<open>bit (concat_bit n k l) m = bit (concat_bit n r s) m\<close> for m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1286
    by (simp add: bit_eq_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1287
  have \<open>take_bit n k = take_bit n r\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1288
  proof (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1289
    fix m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1290
    from * [of m]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1291
    show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1292
      by (auto simp add: bit_take_bit_iff bit_concat_bit_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1293
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1294
  moreover have \<open>push_bit n l = push_bit n s\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1295
  proof (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1296
    fix m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1297
    from * [of m]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1298
    show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1299
      by (auto simp add: bit_push_bit_iff bit_concat_bit_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1300
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1301
  then have \<open>l = s\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1302
    by (simp add: push_bit_eq_mult)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1303
  ultimately show ?Q
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1304
    by (simp add: concat_bit_def)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1305
qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1306
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1307
lemma take_bit_concat_bit_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1308
  \<open>take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1309
  by (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1310
    (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def)  
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1311
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1312
lemma concat_bit_take_bit_eq:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1313
  \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1314
  by (simp add: concat_bit_def [abs_def])
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1315
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1316
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1317
subsection \<open>Taking bits with sign propagation\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1318
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1319
context ring_bit_operations
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1320
begin
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1321
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1322
definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1323
  where \<open>signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\<close>
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1324
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1325
lemma signed_take_bit_eq_if_positive:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1326
  \<open>signed_take_bit n a = take_bit n a\<close> if \<open>\<not> bit a n\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1327
  using that by (simp add: signed_take_bit_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1328
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1329
lemma signed_take_bit_eq_if_negative:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1330
  \<open>signed_take_bit n a = take_bit n a OR NOT (mask n)\<close> if \<open>bit a n\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1331
  using that by (simp add: signed_take_bit_def)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1332
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1333
lemma even_signed_take_bit_iff:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1334
  \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1335
  by (auto simp add: signed_take_bit_def even_or_iff even_mask_iff bit_double_iff)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1336
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1337
lemma bit_signed_take_bit_iff [bit_simps]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1338
  \<open>bit (signed_take_bit m a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> bit a (min m n)\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1339
  by (simp add: signed_take_bit_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff min_def not_le)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1340
    (use exp_eq_0_imp_not_bit in blast)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1341
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1342
lemma signed_take_bit_0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1343
  \<open>signed_take_bit 0 a = - (a mod 2)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1344
  by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1345
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1346
lemma signed_take_bit_Suc:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1347
  \<open>signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1348
proof (rule bit_eqI)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1349
  fix m
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1350
  assume *: \<open>2 ^ m \<noteq> 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1351
  show \<open>bit (signed_take_bit (Suc n) a) m \<longleftrightarrow>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1352
    bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1353
  proof (cases m)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1354
    case 0
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1355
    then show ?thesis
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1356
      by (simp add: even_signed_take_bit_iff)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1357
  next
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1358
    case (Suc m)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1359
    with * have \<open>2 ^ m \<noteq> 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1360
      by (metis mult_not_zero power_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1361
    with Suc show ?thesis
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1362
      by (simp add: bit_signed_take_bit_iff mod2_eq_if bit_double_iff even_bit_succ_iff
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1363
        ac_simps flip: bit_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1364
  qed
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1365
qed
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1366
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1367
lemma signed_take_bit_of_0 [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1368
  \<open>signed_take_bit n 0 = 0\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1369
  by (simp add: signed_take_bit_def)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1370
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1371
lemma signed_take_bit_of_minus_1 [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1372
  \<open>signed_take_bit n (- 1) = - 1\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1373
  by (simp add: signed_take_bit_def take_bit_minus_one_eq_mask mask_eq_exp_minus_1)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1374
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1375
lemma signed_take_bit_Suc_1 [simp]:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1376
  \<open>signed_take_bit (Suc n) 1 = 1\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1377
  by (simp add: signed_take_bit_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1378
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1379
lemma signed_take_bit_rec:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1380
  \<open>signed_take_bit n a = (if n = 0 then - (a mod 2) else a mod 2 + 2 * signed_take_bit (n - 1) (a div 2))\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1381
  by (cases n) (simp_all add: signed_take_bit_Suc)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1382
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1383
lemma signed_take_bit_eq_iff_take_bit_eq:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1384
  \<open>signed_take_bit n a = signed_take_bit n b \<longleftrightarrow> take_bit (Suc n) a = take_bit (Suc n) b\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1385
proof -
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1386
  have \<open>bit (signed_take_bit n a) = bit (signed_take_bit n b) \<longleftrightarrow> bit (take_bit (Suc n) a) = bit (take_bit (Suc n) b)\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1387
    by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1388
      (use exp_eq_0_imp_not_bit in fastforce)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1389
  then show ?thesis
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1390
    by (simp add: bit_eq_iff fun_eq_iff)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1391
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1392
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1393
lemma signed_take_bit_signed_take_bit [simp]:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1394
  \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1395
proof (rule bit_eqI)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1396
  fix q
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1397
  show \<open>bit (signed_take_bit m (signed_take_bit n a)) q \<longleftrightarrow>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1398
    bit (signed_take_bit (min m n) a) q\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1399
    by (simp add: bit_signed_take_bit_iff min_def bit_or_iff bit_not_iff bit_mask_iff bit_take_bit_iff)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1400
      (use le_Suc_ex exp_add_not_zero_imp in blast)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1401
qed
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1402
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1403
lemma signed_take_bit_take_bit:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1404
  \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1405
  by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1406
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1407
lemma take_bit_signed_take_bit:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1408
  \<open>take_bit m (signed_take_bit n a) = take_bit m a\<close> if \<open>m \<le> Suc n\<close>
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1409
  using that by (rule le_SucE; intro bit_eqI)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1410
   (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1411
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1412
end
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1413
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1414
text \<open>Modulus centered around 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1415
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1416
lemma signed_take_bit_eq_concat_bit:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1417
  \<open>signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1418
  by (simp add: concat_bit_def signed_take_bit_def push_bit_minus_one_eq_not_mask)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1419
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1420
lemma signed_take_bit_add:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1421
  \<open>signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1422
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1423
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1424
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1425
     (take_bit (Suc n) (signed_take_bit n k) +
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1426
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1427
    take_bit (Suc n) (k + l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1428
    by (simp add: take_bit_signed_take_bit take_bit_add)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1429
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1430
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1431
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1432
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1433
lemma signed_take_bit_diff:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1434
  \<open>signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1435
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1436
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1437
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1438
     (take_bit (Suc n) (signed_take_bit n k) -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1439
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1440
    take_bit (Suc n) (k - l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1441
    by (simp add: take_bit_signed_take_bit take_bit_diff)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1442
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1443
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1444
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1445
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1446
lemma signed_take_bit_minus:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1447
  \<open>signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1448
  for k :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1449
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1450
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1451
     (- take_bit (Suc n) (signed_take_bit n k)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1452
    take_bit (Suc n) (- k)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1453
    by (simp add: take_bit_signed_take_bit take_bit_minus)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1454
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1455
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1456
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1457
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1458
lemma signed_take_bit_mult:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1459
  \<open>signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1460
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1461
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1462
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1463
     (take_bit (Suc n) (signed_take_bit n k) *
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1464
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1465
    take_bit (Suc n) (k * l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1466
    by (simp add: take_bit_signed_take_bit take_bit_mult)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1467
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1468
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1469
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  1470
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1471
lemma signed_take_bit_eq_take_bit_minus:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1472
  \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1473
  for k :: int
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1474
proof (cases \<open>bit k n\<close>)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1475
  case True
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1476
  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
  1477
    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
  1478
  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
  1479
    by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1480
  with True show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1481
    by (simp flip: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1482
next
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1483
  case False
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1484
  show ?thesis
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1485
    by (rule bit_eqI) (simp add: False bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1486
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1487
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1488
lemma signed_take_bit_eq_take_bit_shift:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1489
  \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1490
  for k :: int
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1491
proof -
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1492
  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
  1493
    by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1494
  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
  1495
    by (simp add: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1496
  also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1497
    by (rule disjunctive_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1498
      (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1499
  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
  1500
  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
  1501
    by (simp only: take_bit_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1502
  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
  1503
    by (simp add: take_bit_Suc_from_most)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1504
  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
  1505
    by (simp add: ac_simps)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1506
  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
  1507
    by (rule disjunctive_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1508
      (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1509
  finally show ?thesis
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1510
    using * ** by (simp add: signed_take_bit_def concat_bit_Suc min_def ac_simps)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1511
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1512
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1513
lemma signed_take_bit_nonnegative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1514
  \<open>0 \<le> signed_take_bit n k \<longleftrightarrow> \<not> bit k n\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1515
  for k :: int
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1516
  by (simp add: signed_take_bit_def not_less concat_bit_def)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1517
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1518
lemma signed_take_bit_negative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1519
  \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1520
  for k :: int
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1521
  by (simp add: signed_take_bit_def not_less concat_bit_def)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1522
73868
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1523
lemma signed_take_bit_int_greater_eq_minus_exp [simp]:
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1524
  \<open>- (2 ^ n) \<le> signed_take_bit n k\<close>
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1525
  for k :: int
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1526
  by (simp add: signed_take_bit_eq_take_bit_shift)
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1527
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1528
lemma signed_take_bit_int_less_exp [simp]:
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1529
  \<open>signed_take_bit n k < 2 ^ n\<close>
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1530
  for k :: int
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1531
  using take_bit_int_less_exp [of \<open>Suc n\<close>]
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1532
  by (simp add: signed_take_bit_eq_take_bit_shift)
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  1533
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1534
lemma signed_take_bit_int_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1535
  \<open>signed_take_bit n k = k \<longleftrightarrow> - (2 ^ n) \<le> k \<and> k < 2 ^ n\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1536
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1537
  by (auto simp add: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps)
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1538
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1539
lemma signed_take_bit_int_eq_self:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1540
  \<open>signed_take_bit n k = k\<close> if \<open>- (2 ^ n) \<le> k\<close> \<open>k < 2 ^ n\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1541
  for k :: int
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1542
  using that by (simp add: signed_take_bit_int_eq_self_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1543
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1544
lemma signed_take_bit_int_less_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1545
  \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1546
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1547
  by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_eq_self_iff algebra_simps)
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1548
    linarith
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1549
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1550
lemma signed_take_bit_int_less_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1551
  \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1552
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1553
  by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_self_iff algebra_simps)
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1554
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1555
lemma signed_take_bit_int_greater_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1556
  \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1557
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1558
  by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_self_iff algebra_simps)
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1559
    linarith
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1560
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1561
lemma signed_take_bit_int_greater_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1562
  \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1563
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1564
  by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_eq_self_iff algebra_simps)
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1565
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1566
lemma signed_take_bit_int_greater_eq:
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1567
  \<open>k + 2 ^ Suc n \<le> signed_take_bit n k\<close> if \<open>k < - (2 ^ n)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1568
  for k :: int
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1569
  using that take_bit_int_greater_eq [of \<open>k + 2 ^ n\<close> \<open>Suc n\<close>]
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1570
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1571
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  1572
lemma signed_take_bit_int_less_eq:
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1573
  \<open>signed_take_bit n k \<le> k - 2 ^ Suc n\<close> if \<open>k \<ge> 2 ^ n\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1574
  for k :: int
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1575
  using that take_bit_int_less_eq [of \<open>Suc n\<close> \<open>k + 2 ^ n\<close>]
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1576
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1577
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1578
lemma signed_take_bit_Suc_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1579
  \<open>signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1580
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1581
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1582
lemma signed_take_bit_Suc_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1583
  \<open>signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1584
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1585
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1586
lemma signed_take_bit_Suc_minus_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1587
  \<open>signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1588
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1589
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1590
lemma signed_take_bit_Suc_minus_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1591
  \<open>signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1592
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1593
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1594
lemma signed_take_bit_numeral_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1595
  \<open>signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1596
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1597
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1598
lemma signed_take_bit_numeral_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1599
  \<open>signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1600
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1601
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1602
lemma signed_take_bit_numeral_minus_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1603
  \<open>signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1604
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1605
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1606
lemma signed_take_bit_numeral_minus_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1607
  \<open>signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1608
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1609
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1610
lemma signed_take_bit_code [code]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1611
  \<open>signed_take_bit n a =
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1612
  (let l = take_bit (Suc n) a
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1613
   in if bit l n then l + push_bit (Suc n) (- 1) else l)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1614
proof -
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1615
  have *: \<open>take_bit (Suc n) a + push_bit n (- 2) =
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1616
    take_bit (Suc n) a OR NOT (mask (Suc n))\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1617
    by (auto simp add: bit_take_bit_iff bit_push_bit_iff bit_not_iff bit_mask_iff disjunctive_add
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1618
       simp flip: push_bit_minus_one_eq_not_mask)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1619
  show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1620
    by (rule bit_eqI)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1621
      (auto simp add: Let_def * bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq bit_not_iff bit_mask_iff bit_or_iff)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1622
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1623
72512
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1624
lemma not_minus_numeral_inc_eq:
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1625
  \<open>NOT (- numeral (Num.inc n)) = (numeral n :: int)\<close>
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1626
  by (simp add: not_int_def sub_inc_One_eq)
83b5911c0164 more lemmas
haftmann
parents: 72508
diff changeset
  1627
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1628
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1629
subsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1630
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1631
instantiation nat :: semiring_bit_operations
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1632
begin
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1633
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1634
definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1635
  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
  1636
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1637
definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1638
  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
  1639
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1640
definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1641
  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
  1642
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1643
definition mask_nat :: \<open>nat \<Rightarrow> nat\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1644
  where \<open>mask n = (2 :: nat) ^ n - 1\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1645
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1646
definition set_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1647
  where \<open>set_bit m n = n OR push_bit m 1\<close> for m n :: nat
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1648
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1649
definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1650
  where \<open>unset_bit m n = (if bit n m then n - push_bit m 1 else n)\<close> for m n :: nat
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1651
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1652
definition flip_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1653
  where \<open>flip_bit m n = n XOR push_bit m 1\<close> for m n :: nat
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1654
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1655
instance proof
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1656
  fix m n q :: nat
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1657
  show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1658
    by (simp add: and_nat_def bit_simps)
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1659
  show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1660
    by (simp add: or_nat_def bit_simps)
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1661
  show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1662
    by (simp add: xor_nat_def bit_simps)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1663
  show \<open>bit (unset_bit m n) q \<longleftrightarrow> bit n q \<and> m \<noteq> q\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1664
  proof (cases \<open>bit n m\<close>)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1665
    case False
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1666
    then show ?thesis by (auto simp add: unset_bit_nat_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1667
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1668
    case True
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1669
    have \<open>push_bit m (drop_bit m n) + take_bit m n = n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1670
      by (fact bits_ident)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1671
    also from \<open>bit n m\<close> have \<open>drop_bit m n = 2 * drop_bit (Suc m) n  + 1\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1672
      by (simp add: drop_bit_Suc drop_bit_half even_drop_bit_iff_not_bit ac_simps)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1673
    finally have \<open>push_bit m (2 * drop_bit (Suc m) n) + take_bit m n + push_bit m 1 = n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1674
      by (simp only: push_bit_add ac_simps)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1675
    then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) + take_bit m n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1676
      by simp
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1677
    then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) OR take_bit m n\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1678
      by (simp add: or_nat_def bit_simps flip: disjunctive_add)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1679
    with \<open>bit n m\<close> show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1680
      by (auto simp add: unset_bit_nat_def or_nat_def bit_simps)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1681
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1682
qed (simp_all add: mask_nat_def set_bit_nat_def flip_bit_nat_def)
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1683
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1684
end
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1685
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1686
lemma and_nat_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1687
  \<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
  1688
  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
  1689
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1690
lemma or_nat_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1691
  \<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
  1692
  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
  1693
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1694
lemma xor_nat_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1695
  \<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
  1696
  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
  1697
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1698
lemma Suc_0_and_eq [simp]:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1699
  \<open>Suc 0 AND n = n mod 2\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1700
  using one_and_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1701
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1702
lemma and_Suc_0_eq [simp]:
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1703
  \<open>n AND Suc 0 = n mod 2\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1704
  using and_one_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1705
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1706
lemma Suc_0_or_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1707
  \<open>Suc 0 OR n = n + of_bool (even n)\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1708
  using one_or_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1709
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1710
lemma or_Suc_0_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1711
  \<open>n OR Suc 0 = n + of_bool (even n)\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1712
  using or_one_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1713
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1714
lemma Suc_0_xor_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1715
  \<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
  1716
  using one_xor_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1717
71822
67cc2319104f prefer _ mod 2 over of_bool (odd _)
haftmann
parents: 71821
diff changeset
  1718
lemma xor_Suc_0_eq:
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1719
  \<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
  1720
  using xor_one_eq [of n] by simp
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1721
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1722
context semiring_bit_operations
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1723
begin
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1724
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1725
lemma of_nat_and_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1726
  \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1727
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1728
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1729
lemma of_nat_or_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1730
  \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1731
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1732
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1733
lemma of_nat_xor_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1734
  \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1735
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1736
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1737
end
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1738
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1739
context ring_bit_operations
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1740
begin
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1741
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1742
lemma of_nat_mask_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1743
  \<open>of_nat (mask n) = mask n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1744
  by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1745
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1746
end
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  1747
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1748
lemma Suc_mask_eq_exp:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1749
  \<open>Suc (mask n) = 2 ^ n\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1750
  by (simp add: mask_eq_exp_minus_1)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1751
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1752
lemma less_eq_mask:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1753
  \<open>n \<le> mask n\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1754
  by (simp add: mask_eq_exp_minus_1 le_diff_conv2)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1755
    (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1756
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1757
lemma less_mask:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1758
  \<open>n < mask n\<close> if \<open>Suc 0 < n\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1759
proof -
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1760
  define m where \<open>m = n - 2\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1761
  with that have *: \<open>n = m + 2\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1762
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1763
  have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1764
    by (induction m) simp_all
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1765
  then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1766
    by (simp add: Suc_mask_eq_exp)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1767
  then have \<open>m + 2 < mask (m + 2)\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1768
    by (simp add: less_le)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1769
  with * show ?thesis
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1770
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1771
qed
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1772
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1773
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1774
subsection \<open>Instances for \<^typ>\<open>integer\<close> and \<^typ>\<open>natural\<close>\<close>
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1775
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1776
unbundle integer.lifting natural.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1777
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1778
instantiation integer :: ring_bit_operations
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1779
begin
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1780
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1781
lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1782
  is not .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1783
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1784
lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1785
  is \<open>and\<close> .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1786
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1787
lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1788
  is or .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1789
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1790
lift_definition xor_integer ::  \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1791
  is xor .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1792
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1793
lift_definition mask_integer :: \<open>nat \<Rightarrow> integer\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1794
  is mask .
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1795
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1796
lift_definition set_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1797
  is set_bit .
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1798
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1799
lift_definition unset_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1800
  is unset_bit .
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1801
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1802
lift_definition flip_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1803
  is flip_bit .
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1804
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1805
instance by (standard; transfer)
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1806
  (simp_all add: minus_eq_not_minus_1 mask_eq_exp_minus_1
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1807
    bit_not_iff bit_and_iff bit_or_iff bit_xor_iff
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1808
    set_bit_def bit_unset_bit_iff flip_bit_def)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1809
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1810
end
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1811
72083
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1812
lemma [code]:
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1813
  \<open>mask n = 2 ^ n - (1::integer)\<close>
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1814
  by (simp add: mask_eq_exp_minus_1)
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1815
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1816
instantiation natural :: semiring_bit_operations
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1817
begin
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1818
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1819
lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1820
  is \<open>and\<close> .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1821
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1822
lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1823
  is or .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1824
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1825
lift_definition xor_natural ::  \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1826
  is xor .
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1827
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1828
lift_definition mask_natural :: \<open>nat \<Rightarrow> natural\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1829
  is mask .
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1830
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1831
lift_definition set_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1832
  is set_bit .
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1833
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1834
lift_definition unset_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1835
  is unset_bit .
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1836
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1837
lift_definition flip_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1838
  is flip_bit .
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1839
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1840
instance by (standard; transfer)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1841
  (simp_all add: mask_eq_exp_minus_1
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1842
    bit_and_iff bit_or_iff bit_xor_iff
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1843
    set_bit_def bit_unset_bit_iff flip_bit_def)
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1844
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1845
end
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1846
72083
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1847
lemma [code]:
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1848
  \<open>integer_of_natural (mask n) = mask n\<close>
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1849
  by transfer (simp add: mask_eq_exp_minus_1 of_nat_diff)
3ec876181527 further refinement of code equations for mask operation
haftmann
parents: 72082
diff changeset
  1850
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1851
lifting_update integer.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1852
lifting_forget integer.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1853
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1854
lifting_update natural.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1855
lifting_forget natural.lifting
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1856
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1857
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1858
subsection \<open>Key ideas of bit operations\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1859
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1860
text \<open>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1861
  When formalizing bit operations, it is tempting to represent
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1862
  bit values as explicit lists over a binary type. This however
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1863
  is a bad idea, mainly due to the inherent ambiguities in
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1864
  representation concerning repeating leading bits.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1865
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1866
  Hence this approach avoids such explicit lists altogether
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1867
  following an algebraic path:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1868
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1869
  \<^item> Bit values are represented by numeric types: idealized
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1870
    unbounded bit values can be represented by type \<^typ>\<open>int\<close>,
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1871
    bounded bit values by quotient types over \<^typ>\<open>int\<close>.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1872
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1873
  \<^item> (A special case are idealized unbounded bit values ending
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1874
    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
  1875
    only support a restricted set of operations).
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1876
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1877
  \<^item> From this idea follows that
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1878
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1879
      \<^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
  1880
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1881
      \<^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
  1882
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1883
  \<^item> Concerning bounded bit values, iterated shifts to the left
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1884
    may result in eliminating all bits by shifting them all
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1885
    beyond the boundary.  The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1886
    represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1887
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
  1888
  \<^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
  1889
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1890
  \<^item> This leads to the most fundamental properties of bit values:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1891
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1892
      \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1893
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1894
      \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1895
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1896
  \<^item> Typical operations are characterized as follows:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1897
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1898
      \<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1899
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1900
      \<^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
  1901
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1902
      \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1903
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1904
      \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1905
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1906
      \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1907
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1908
      \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1909
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1910
      \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1911
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1912
      \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1913
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1914
      \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1915
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1916
      \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1917
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1918
      \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1919
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1920
      \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]}
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1921
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1922
      \<^item> Signed truncation, or modulus centered around \<^term>\<open>0::int\<close>: @{thm signed_take_bit_def [no_vars]}
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1923
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1924
      \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]}
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1925
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1926
      \<^item> (Bounded) conversion from and to a list of bits: @{thm horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]}
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1927
\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  1928
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1929
code_identifier
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1930
  type_class semiring_bits \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1931
  (SML) Bit_Operations.semiring_bits and (OCaml) Bit_Operations.semiring_bits and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bits
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1932
| class_relation semiring_bits < semiring_parity \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1933
  (SML) Bit_Operations.semiring_parity_semiring_bits and (OCaml) Bit_Operations.semiring_parity_semiring_bits and (Haskell) Bit_Operations.semiring_parity_semiring_bits and (Scala) Bit_Operations.semiring_parity_semiring_bits
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1934
| constant bit \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1935
  (SML) Bit_Operations.bit and (OCaml) Bit_Operations.bit and (Haskell) Bit_Operations.bit and (Scala) Bit_Operations.bit
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1936
| class_instance nat :: semiring_bits \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1937
  (SML) Bit_Operations.semiring_bits_nat and (OCaml) Bit_Operations.semiring_bits_nat and (Haskell) Bit_Operations.semiring_bits_nat and (Scala) Bit_Operations.semiring_bits_nat
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1938
| class_instance int :: semiring_bits \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1939
  (SML) Bit_Operations.semiring_bits_int and (OCaml) Bit_Operations.semiring_bits_int and (Haskell) Bit_Operations.semiring_bits_int and (Scala) Bit_Operations.semiring_bits_int
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1940
| type_class semiring_bit_shifts \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1941
  (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bit_shifts
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1942
| class_relation semiring_bit_shifts < semiring_bits \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1943
  (SML) Bit_Operations.semiring_bits_semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bits_semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits_semiring_bit_shifts and (Scala) Bit_Operations.semiring_bits_semiring_bit_shifts
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1944
| constant push_bit \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1945
  (SML) Bit_Operations.push_bit and (OCaml) Bit_Operations.push_bit and (Haskell) Bit_Operations.push_bit and (Scala) Bit_Operations.push_bit
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1946
| constant drop_bit \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1947
  (SML) Bit_Operations.drop_bit and (OCaml) Bit_Operations.drop_bit and (Haskell) Bit_Operations.drop_bit and (Scala) Bit_Operations.drop_bit
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1948
| constant take_bit \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1949
  (SML) Bit_Operations.take_bit and (OCaml) Bit_Operations.take_bit and (Haskell) Bit_Operations.take_bit and (Scala) Bit_Operations.take_bit
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1950
| class_instance nat :: semiring_bit_shifts \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1951
  (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1952
| class_instance int :: semiring_bit_shifts \<rightharpoonup>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1953
  (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1954
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  1955
end