src/HOL/Bit_Operations.thy
author haftmann
Tue, 03 Aug 2021 13:53:22 +0000
changeset 74108 3146646a43a7
parent 74101 d804e93ae9ff
child 74123 7c5842b06114
permissions -rw-r--r--
simplified hierarchy of type classes for bit operations
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
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
     7
  imports Presburger Groups_List
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
     8
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
     9
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    10
subsection \<open>Abstract bit structures\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    11
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    12
class semiring_bits = semiring_parity +
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    13
  assumes bits_induct [case_names stable rec]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    14
    \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    15
     \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a))
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    16
        \<Longrightarrow> P a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    17
  assumes bits_div_0 [simp]: \<open>0 div a = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    18
    and bits_div_by_1 [simp]: \<open>a div 1 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    19
    and bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    20
    and even_succ_div_2 [simp]: \<open>even a \<Longrightarrow> (1 + a) div 2 = a div 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    21
    and even_mask_div_iff: \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    22
    and exp_div_exp_eq: \<open>2 ^ m div 2 ^ n = of_bool (2 ^ m \<noteq> 0 \<and> m \<ge> n) * 2 ^ (m - n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    23
    and div_exp_eq: \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    24
    and mod_exp_eq: \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    25
    and mult_exp_mod_exp_eq: \<open>m \<le> n \<Longrightarrow> (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    26
    and div_exp_mod_exp_eq: \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    27
    and even_mult_exp_div_exp_iff: \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> m > n \<or> 2 ^ n = 0 \<or> (m \<le> n \<and> even (a div 2 ^ (n - m)))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    28
  fixes bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    29
  assumes bit_iff_odd: \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    30
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    31
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    32
text \<open>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    33
  Having \<^const>\<open>bit\<close> as definitional class operation
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    34
  takes into account that specific instances can be implemented
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    35
  differently wrt. code generation.
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    36
\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    37
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    38
lemma bits_div_by_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    39
  \<open>a div 0 = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    40
  by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    41
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    42
lemma bits_1_div_2 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    43
  \<open>1 div 2 = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    44
  using even_succ_div_2 [of 0] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    45
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    46
lemma bits_1_div_exp [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    47
  \<open>1 div 2 ^ n = of_bool (n = 0)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    48
  using div_exp_eq [of 1 1] by (cases n) simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    49
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    50
lemma even_succ_div_exp [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    51
  \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    52
proof (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    53
  case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    54
  with that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    55
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    56
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    57
  case (Suc n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    58
  with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    59
  proof (induction n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    60
    case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    61
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    62
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    63
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    64
    case (Suc n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    65
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    66
      using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    67
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    68
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    69
  with Suc show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    70
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    71
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    72
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    73
lemma even_succ_mod_exp [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    74
  \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    75
  using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] that
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    76
  apply simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    77
  by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    78
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    79
lemma bits_mod_by_1 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    80
  \<open>a mod 1 = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    81
  using div_mult_mod_eq [of a 1] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    82
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    83
lemma bits_mod_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    84
  \<open>0 mod a = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    85
  using div_mult_mod_eq [of 0 a] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    86
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    87
lemma bits_one_mod_two_eq_one [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    88
  \<open>1 mod 2 = 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    89
  by (simp add: mod2_eq_if)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    90
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    91
lemma bit_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    92
  \<open>bit a 0 \<longleftrightarrow> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    93
  by (simp add: bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    94
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    95
lemma bit_Suc:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    96
  \<open>bit a (Suc n) \<longleftrightarrow> bit (a div 2) n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    97
  using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    98
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    99
lemma bit_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   100
  \<open>bit a n \<longleftrightarrow> (if n = 0 then odd a else bit (a div 2) (n - 1))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   101
  by (cases n) (simp_all add: bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   102
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   103
lemma bit_0_eq [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   104
  \<open>bit 0 = bot\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   105
  by (simp add: fun_eq_iff bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   106
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   107
context
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   108
  fixes a
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   109
  assumes stable: \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   110
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   111
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   112
lemma bits_stable_imp_add_self:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   113
  \<open>a + a mod 2 = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   114
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   115
  have \<open>a div 2 * 2 + a mod 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   116
    by (fact div_mult_mod_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   117
  then have \<open>a * 2 + a mod 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   118
    by (simp add: stable)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   119
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   120
    by (simp add: mult_2_right ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   121
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   122
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   123
lemma stable_imp_bit_iff_odd:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   124
  \<open>bit a n \<longleftrightarrow> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   125
  by (induction n) (simp_all add: stable bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   126
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   127
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   128
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   129
lemma bit_iff_idd_imp_stable:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   130
  \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   131
using that proof (induction a rule: bits_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   132
  case (stable a)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   133
  then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   134
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   135
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   136
  case (rec a b)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   137
  from rec.prems [of 1] have [simp]: \<open>b = odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   138
    by (simp add: rec.hyps bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   139
  from rec.hyps have hyp: \<open>(of_bool (odd a) + 2 * a) div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   140
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   141
  have \<open>bit a n \<longleftrightarrow> odd a\<close> for n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   142
    using rec.prems [of \<open>Suc n\<close>] by (simp add: hyp bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   143
  then have \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   144
    by (rule rec.IH)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   145
  then have \<open>of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   146
    by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   147
  also have \<open>\<dots> = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   148
    using mult_div_mod_eq [of 2 a]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   149
    by (simp add: of_bool_odd_eq_mod_2)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   150
  finally show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   151
    using \<open>a div 2 = a\<close> by (simp add: hyp)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   152
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   153
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   154
lemma exp_eq_0_imp_not_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   155
  \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   156
  using that by (simp add: bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   157
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   158
lemma bit_eqI:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   159
  \<open>a = b\<close> if \<open>\<And>n. 2 ^ n \<noteq> 0 \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   160
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   161
  have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   162
  proof (cases \<open>2 ^ n = 0\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   163
    case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   164
    then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   165
      by (simp add: exp_eq_0_imp_not_bit)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   166
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   167
    case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   168
    then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   169
      by (rule that)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   170
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   171
  then show ?thesis proof (induction a arbitrary: b rule: bits_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   172
    case (stable a)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   173
    from stable(2) [of 0] have **: \<open>even b \<longleftrightarrow> even a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   174
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   175
    have \<open>b div 2 = b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   176
    proof (rule bit_iff_idd_imp_stable)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   177
      fix n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   178
      from stable have *: \<open>bit b n \<longleftrightarrow> bit a n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   179
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   180
      also have \<open>bit a n \<longleftrightarrow> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   181
        using stable by (simp add: stable_imp_bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   182
      finally show \<open>bit b n \<longleftrightarrow> odd b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   183
        by (simp add: **)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   184
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   185
    from ** have \<open>a mod 2 = b mod 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   186
      by (simp add: mod2_eq_if)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   187
    then have \<open>a mod 2 + (a + b) = b mod 2 + (a + b)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   188
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   189
    then have \<open>a + a mod 2 + b = b + b mod 2 + a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   190
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   191
    with \<open>a div 2 = a\<close> \<open>b div 2 = b\<close> show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   192
      by (simp add: bits_stable_imp_add_self)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   193
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   194
    case (rec a p)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   195
    from rec.prems [of 0] have [simp]: \<open>p = odd b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   196
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   197
    from rec.hyps have \<open>bit a n \<longleftrightarrow> bit (b div 2) n\<close> for n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   198
      using rec.prems [of \<open>Suc n\<close>] by (simp add: bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   199
    then have \<open>a = b div 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   200
      by (rule rec.IH)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   201
    then have \<open>2 * a = 2 * (b div 2)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   202
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   203
    then have \<open>b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   204
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   205
    also have \<open>\<dots> = b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   206
      by (fact mod_mult_div_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   207
    finally show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   208
      by (auto simp add: mod2_eq_if)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   209
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   210
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   211
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   212
lemma bit_eq_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   213
  \<open>a = b \<longleftrightarrow> (\<forall>n. bit a n \<longleftrightarrow> bit b n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   214
  by (auto intro: bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   215
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   216
named_theorems bit_simps \<open>Simplification rules for \<^const>\<open>bit\<close>\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   217
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   218
lemma bit_exp_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   219
  \<open>bit (2 ^ m) n \<longleftrightarrow> 2 ^ m \<noteq> 0 \<and> m = n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   220
  by (auto simp add: bit_iff_odd exp_div_exp_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   221
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   222
lemma bit_1_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   223
  \<open>bit 1 n \<longleftrightarrow> 1 \<noteq> 0 \<and> n = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   224
  using bit_exp_iff [of 0 n] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   225
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   226
lemma bit_2_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   227
  \<open>bit 2 n \<longleftrightarrow> 2 \<noteq> 0 \<and> n = 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   228
  using bit_exp_iff [of 1 n] by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   229
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   230
lemma even_bit_succ_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   231
  \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   232
  using that by (cases \<open>n = 0\<close>) (simp_all add: bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   233
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   234
lemma odd_bit_iff_bit_pred:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   235
  \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   236
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   237
  from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> ..
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   238
  moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   239
    using even_bit_succ_iff by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   240
  ultimately show ?thesis by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   241
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   242
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   243
lemma bit_double_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   244
  \<open>bit (2 * a) n \<longleftrightarrow> bit a (n - 1) \<and> n \<noteq> 0 \<and> 2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   245
  using even_mult_exp_div_exp_iff [of a 1 n]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   246
  by (cases n, auto simp add: bit_iff_odd ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   247
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   248
lemma bit_eq_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   249
  \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close> (is \<open>?P = ?Q\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   250
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   251
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   252
  then show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   253
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   254
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   255
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   256
  then have \<open>even a \<longleftrightarrow> even b\<close> and \<open>a div 2 = b div 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   257
    by simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   258
  show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   259
  proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   260
    fix n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   261
    show \<open>bit a n \<longleftrightarrow> bit b n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   262
    proof (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   263
      case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   264
      with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   265
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   266
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   267
      case (Suc n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   268
      moreover from \<open>a div 2 = b div 2\<close> have \<open>bit (a div 2) n = bit (b div 2) n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   269
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   270
      ultimately show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   271
        by (simp add: bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   272
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   273
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   274
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   275
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   276
lemma bit_mod_2_iff [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   277
  \<open>bit (a mod 2) n \<longleftrightarrow> n = 0 \<and> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   278
  by (cases a rule: parity_cases) (simp_all add: bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   279
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   280
lemma bit_mask_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   281
  \<open>bit (2 ^ m - 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   282
  by (simp add: bit_iff_odd even_mask_div_iff not_le)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   283
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   284
lemma bit_Numeral1_iff [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   285
  \<open>bit (numeral Num.One) n \<longleftrightarrow> n = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   286
  by (simp add: bit_rec)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   287
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   288
lemma exp_add_not_zero_imp:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   289
  \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> if \<open>2 ^ (m + n) \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   290
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   291
  have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   292
  proof (rule notI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   293
    assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   294
    then have \<open>2 ^ (m + n) = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   295
      by (rule disjE) (simp_all add: power_add)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   296
    with that show False ..
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   297
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   298
  then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   299
    by simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   300
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   301
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   302
lemma bit_disjunctive_add_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   303
  \<open>bit (a + b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   304
  if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   305
proof (cases \<open>2 ^ n = 0\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   306
  case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   307
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   308
    by (simp add: exp_eq_0_imp_not_bit)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   309
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   310
  case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   311
  with that show ?thesis proof (induction n arbitrary: a b)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   312
    case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   313
    from "0.prems"(1) [of 0] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   314
      by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   315
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   316
    case (Suc n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   317
    from Suc.prems(1) [of 0] have even: \<open>even a \<or> even b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   318
      by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   319
    have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   320
      using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   321
    from Suc.prems(2) have \<open>2 * 2 ^ n \<noteq> 0\<close> \<open>2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   322
      by (auto simp add: mult_2)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   323
    have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   324
      using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   325
    also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   326
      using even by (auto simp add: algebra_simps mod2_eq_if)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   327
    finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   328
      using \<open>2 * 2 ^ n \<noteq> 0\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   329
    also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   330
      using bit \<open>2 ^ n \<noteq> 0\<close> by (rule Suc.IH)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   331
    finally show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   332
      by (simp add: bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   333
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   334
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   335
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   336
lemma
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   337
  exp_add_not_zero_imp_left: \<open>2 ^ m \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   338
  and exp_add_not_zero_imp_right: \<open>2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   339
  if \<open>2 ^ (m + n) \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   340
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   341
  have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   342
  proof (rule notI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   343
    assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   344
    then have \<open>2 ^ (m + n) = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   345
      by (rule disjE) (simp_all add: power_add)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   346
    with that show False ..
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   347
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   348
  then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   349
    by simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   350
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   351
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   352
lemma exp_not_zero_imp_exp_diff_not_zero:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   353
  \<open>2 ^ (n - m) \<noteq> 0\<close> if \<open>2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   354
proof (cases \<open>m \<le> n\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   355
  case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   356
  moreover define q where \<open>q = n - m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   357
  ultimately have \<open>n = m + q\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   358
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   359
  with that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   360
    by (simp add: exp_add_not_zero_imp_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   361
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   362
  case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   363
  with that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   364
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   365
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   366
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   367
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   368
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   369
lemma nat_bit_induct [case_names zero even odd]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   370
  "P n" if zero: "P 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   371
    and even: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   372
    and odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   373
proof (induction n rule: less_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   374
  case (less n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   375
  show "P n"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   376
  proof (cases "n = 0")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   377
    case True with zero show ?thesis by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   378
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   379
    case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   380
    with less have hyp: "P (n div 2)" by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   381
    show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   382
    proof (cases "even n")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   383
      case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   384
      then have "n \<noteq> 1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   385
        by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   386
      with \<open>n \<noteq> 0\<close> have "n div 2 > 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   387
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   388
      with \<open>even n\<close> hyp even [of "n div 2"] show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   389
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   390
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   391
      case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   392
      with hyp odd [of "n div 2"] show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   393
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   394
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   395
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   396
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   397
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   398
instantiation nat :: semiring_bits
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   399
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   400
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   401
definition bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   402
  where \<open>bit_nat m n \<longleftrightarrow> odd (m div 2 ^ n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   403
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   404
instance
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   405
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   406
  show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   407
    and rec: \<open>\<And>n b. P n \<Longrightarrow> (of_bool b + 2 * n) div 2 = n \<Longrightarrow> P (of_bool b + 2 * n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   408
    for P and n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   409
  proof (induction n rule: nat_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   410
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   411
    from stable [of 0] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   412
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   413
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   414
    case (even n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   415
    with rec [of n False] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   416
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   417
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   418
    case (odd n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   419
    with rec [of n True] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   420
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   421
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   422
  show \<open>q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   423
    for q m n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   424
    apply (auto simp add: less_iff_Suc_add power_add mod_mod_cancel split: split_min_lin)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   425
    apply (metis div_mult2_eq mod_div_trivial mod_eq_self_iff_div_eq_0 mod_mult_self2_is_0 power_commutes)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   426
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   427
  show \<open>(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\<close> if \<open>m \<le> n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   428
    for q m n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   429
    using that
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   430
    apply (auto simp add: mod_mod_cancel div_mult2_eq power_add mod_mult2_eq le_iff_add split: split_min_lin)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   431
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   432
  show \<open>even ((2 ^ m - (1::nat)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::nat) \<or> m \<le> n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   433
    for m n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   434
    using even_mask_div_iff' [where ?'a = nat, of m n] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   435
  show \<open>even (q * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::nat) ^ n = 0 \<or> m \<le> n \<and> even (q div 2 ^ (n - m))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   436
    for m n q r :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   437
    apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   438
    apply (metis (full_types) dvd_mult dvd_mult_imp_div dvd_power_iff_le not_less not_less_eq order_refl power_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   439
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   440
qed (auto simp add: div_mult2_eq mod_mult2_eq power_add power_diff bit_nat_def)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   441
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   442
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   443
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   444
lemma int_bit_induct [case_names zero minus even odd]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   445
  "P k" if zero_int: "P 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   446
    and minus_int: "P (- 1)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   447
    and even_int: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   448
    and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   449
proof (cases "k \<ge> 0")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   450
  case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   451
  define n where "n = nat k"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   452
  with True have "k = int n"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   453
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   454
  then show "P k"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   455
  proof (induction n arbitrary: k rule: nat_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   456
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   457
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   458
      by (simp add: zero_int)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   459
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   460
    case (even n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   461
    have "P (int n * 2)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   462
      by (rule even_int) (use even in simp_all)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   463
    with even show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   464
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   465
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   466
    case (odd n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   467
    have "P (1 + (int n * 2))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   468
      by (rule odd_int) (use odd in simp_all)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   469
    with odd show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   470
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   471
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   472
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   473
  case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   474
  define n where "n = nat (- k - 1)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   475
  with False have "k = - int n - 1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   476
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   477
  then show "P k"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   478
  proof (induction n arbitrary: k rule: nat_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   479
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   480
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   481
      by (simp add: minus_int)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   482
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   483
    case (even n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   484
    have "P (1 + (- int (Suc n) * 2))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   485
      by (rule odd_int) (use even in \<open>simp_all add: algebra_simps\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   486
    also have "\<dots> = - int (2 * n) - 1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   487
      by (simp add: algebra_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   488
    finally show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   489
      using even.prems by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   490
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   491
    case (odd n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   492
    have "P (- int (Suc n) * 2)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   493
      by (rule even_int) (use odd in \<open>simp_all add: algebra_simps\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   494
    also have "\<dots> = - int (Suc (2 * n)) - 1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   495
      by (simp add: algebra_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   496
    finally show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   497
      using odd.prems by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   498
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   499
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   500
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   501
context semiring_bits
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   502
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   503
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   504
lemma bit_of_bool_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   505
  \<open>bit (of_bool b) n \<longleftrightarrow> b \<and> n = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   506
  by (simp add: bit_1_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   507
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   508
lemma even_of_nat_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   509
  \<open>even (of_nat n) \<longleftrightarrow> even n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   510
  by (induction n rule: nat_bit_induct) simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   511
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   512
lemma bit_of_nat_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   513
  \<open>bit (of_nat m) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit m n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   514
proof (cases \<open>(2::'a) ^ n = 0\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   515
  case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   516
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   517
    by (simp add: exp_eq_0_imp_not_bit)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   518
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   519
  case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   520
  then have \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   521
  proof (induction m arbitrary: n rule: nat_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   522
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   523
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   524
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   525
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   526
    case (even m)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   527
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   528
      by (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   529
        (auto simp add: bit_double_iff Bit_Operations.bit_double_iff dest: mult_not_zero)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   530
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   531
    case (odd m)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   532
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   533
      by (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   534
         (auto simp add: bit_double_iff even_bit_succ_iff Bit_Operations.bit_Suc dest: mult_not_zero)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   535
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   536
  with False show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   537
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   538
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   539
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   540
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   541
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   542
instantiation int :: semiring_bits
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   543
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   544
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   545
definition bit_int :: \<open>int \<Rightarrow> nat \<Rightarrow> bool\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   546
  where \<open>bit_int k n \<longleftrightarrow> odd (k div 2 ^ n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   547
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   548
instance
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   549
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   550
  show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   551
    and rec: \<open>\<And>k b. P k \<Longrightarrow> (of_bool b + 2 * k) div 2 = k \<Longrightarrow> P (of_bool b + 2 * k)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   552
    for P and k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   553
  proof (induction k rule: int_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   554
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   555
    from stable [of 0] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   556
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   557
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   558
    case minus
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   559
    from stable [of \<open>- 1\<close>] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   560
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   561
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   562
    case (even k)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   563
    with rec [of k False] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   564
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   565
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   566
    case (odd k)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   567
    with rec [of k True] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   568
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   569
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   570
  show \<open>(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   571
    for m n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   572
  proof (cases \<open>m < n\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   573
    case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   574
    then have \<open>n = m + (n - m)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   575
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   576
    then have \<open>(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   577
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   578
    also have \<open>\<dots> = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   579
      by (simp add: power_add)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   580
    also have \<open>\<dots> = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   581
      by (simp add: zdiv_zmult2_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   582
    finally show ?thesis using \<open>m < n\<close> by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   583
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   584
    case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   585
    then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   586
      by (simp add: power_diff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   587
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   588
  show \<open>k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   589
    for m n :: nat and k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   590
    using mod_exp_eq [of \<open>nat k\<close> m n]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   591
    apply (auto simp add: mod_mod_cancel zdiv_zmult2_eq power_add zmod_zmult2_eq le_iff_add split: split_min_lin)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   592
     apply (auto simp add: less_iff_Suc_add mod_mod_cancel power_add)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   593
    apply (simp only: flip: mult.left_commute [of \<open>2 ^ m\<close>])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   594
    apply (subst zmod_zmult2_eq) apply simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   595
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   596
  show \<open>(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   597
    if \<open>m \<le> n\<close> for m n :: nat and k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   598
    using that
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   599
    apply (auto simp add: power_add zmod_zmult2_eq le_iff_add split: split_min_lin)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   600
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   601
  show \<open>even ((2 ^ m - (1::int)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::int) \<or> m \<le> n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   602
    for m n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   603
    using even_mask_div_iff' [where ?'a = int, of m n] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   604
  show \<open>even (k * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::int) ^ n = 0 \<or> m \<le> n \<and> even (k div 2 ^ (n - m))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   605
    for m n :: nat and k l :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   606
    apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   607
    apply (metis Suc_leI dvd_mult dvd_mult_imp_div dvd_power_le dvd_refl power.simps(2))
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   608
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   609
qed (auto simp add: zdiv_zmult2_eq zmod_zmult2_eq power_add power_diff not_le bit_int_def)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   610
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   611
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   612
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   613
lemma bit_not_int_iff':
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   614
  \<open>bit (- k - 1) n \<longleftrightarrow> \<not> bit k n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   615
  for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   616
proof (induction n arbitrary: k)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   617
  case 0
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   618
  show ?case
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   619
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   620
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   621
  case (Suc n)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   622
  have \<open>- k - 1 = - (k + 2) + 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   623
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   624
  also have \<open>(- (k + 2) + 1) div 2 = - (k div 2) - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   625
  proof (cases \<open>even k\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   626
    case True
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   627
    then have \<open>- k div 2 = - (k div 2)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   628
      by rule (simp flip: mult_minus_right)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   629
    with True show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   630
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   631
  next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   632
    case False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   633
    have \<open>4 = 2 * (2::int)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   634
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   635
    also have \<open>2 * 2 div 2 = (2::int)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   636
      by (simp only: nonzero_mult_div_cancel_left)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   637
    finally have *: \<open>4 div 2 = (2::int)\<close> .
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   638
    from False obtain l where k: \<open>k = 2 * l + 1\<close> ..
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   639
    then have \<open>- k - 2 = 2 * - (l + 2) + 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   640
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   641
    then have \<open>(- k - 2) div 2 + 1 = - (k div 2) - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   642
      by (simp flip: mult_minus_right add: *) (simp add: k)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   643
    with False show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   644
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   645
  qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   646
  finally have \<open>(- k - 1) div 2 = - (k div 2) - 1\<close> .
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   647
  with Suc show ?case
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   648
    by (simp add: bit_Suc)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   649
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   650
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   651
lemma bit_nat_iff [bit_simps]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   652
  \<open>bit (nat k) n \<longleftrightarrow> k \<ge> 0 \<and> bit k n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   653
proof (cases \<open>k \<ge> 0\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   654
  case True
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   655
  moreover define m where \<open>m = nat k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   656
  ultimately have \<open>k = int m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   657
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   658
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   659
    by (simp add: bit_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   660
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   661
  case False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   662
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   663
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   664
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   665
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   666
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   667
subsection \<open>Bit operations\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   668
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   669
class semiring_bit_operations = semiring_bits +
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   670
  fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>AND\<close> 64)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   671
    and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>OR\<close> 59)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   672
    and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr \<open>XOR\<close> 59)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   673
    and mask :: \<open>nat \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   674
    and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   675
    and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   676
    and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   677
    and push_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   678
    and drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   679
    and take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   680
  assumes bit_and_iff [bit_simps]: \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   681
    and bit_or_iff [bit_simps]: \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   682
    and bit_xor_iff [bit_simps]: \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   683
    and mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   684
    and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   685
    and bit_unset_bit_iff [bit_simps]: \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   686
    and flip_bit_eq_xor: \<open>flip_bit n a = a XOR push_bit n 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   687
    and push_bit_eq_mult: \<open>push_bit n a = a * 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   688
    and drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   689
    and take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   690
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   691
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   692
text \<open>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   693
  We want the bitwise operations to bind slightly weaker
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   694
  than \<open>+\<close> and \<open>-\<close>.
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   695
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   696
  Logically, \<^const>\<open>push_bit\<close>,
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   697
  \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   698
  as separate operations makes proofs easier, otherwise proof automation
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   699
  would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> in a way obfuscating the basic
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   700
  algebraic relationships between those operations.
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   701
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   702
  For the sake of code generation operations 
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   703
  are specified as definitional class operations,
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   704
  taking into account that specific instances of these can be implemented
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   705
  differently wrt. code generation.
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   706
\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   707
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   708
sublocale "and": semilattice \<open>(AND)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   709
  by standard (auto simp add: bit_eq_iff bit_and_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   710
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   711
sublocale or: semilattice_neutr \<open>(OR)\<close> 0
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   712
  by standard (auto simp add: bit_eq_iff bit_or_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   713
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   714
sublocale xor: comm_monoid \<open>(XOR)\<close> 0
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   715
  by standard (auto simp add: bit_eq_iff bit_xor_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   716
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   717
lemma even_and_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   718
  \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   719
  using bit_and_iff [of a b 0] by auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   720
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   721
lemma even_or_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   722
  \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   723
  using bit_or_iff [of a b 0] by auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   724
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   725
lemma even_xor_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   726
  \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   727
  using bit_xor_iff [of a b 0] by auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   728
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   729
lemma zero_and_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   730
  \<open>0 AND a = 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   731
  by (simp add: bit_eq_iff bit_and_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   732
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   733
lemma and_zero_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   734
  \<open>a AND 0 = 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   735
  by (simp add: bit_eq_iff bit_and_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   736
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   737
lemma one_and_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   738
  \<open>1 AND a = a mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   739
  by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   740
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   741
lemma and_one_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   742
  \<open>a AND 1 = a mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   743
  using one_and_eq [of a] by (simp add: ac_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   744
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   745
lemma one_or_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   746
  \<open>1 OR a = a + of_bool (even a)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   747
  by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   748
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   749
lemma or_one_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   750
  \<open>a OR 1 = a + of_bool (even a)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   751
  using one_or_eq [of a] by (simp add: ac_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   752
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   753
lemma one_xor_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   754
  \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   755
  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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   756
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   757
lemma xor_one_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   758
  \<open>a XOR 1 = a + of_bool (even a) - of_bool (odd a)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   759
  using one_xor_eq [of a] by (simp add: ac_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   760
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   761
lemma bit_iff_odd_drop_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   762
  \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   763
  by (simp add: bit_iff_odd drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   764
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   765
lemma even_drop_bit_iff_not_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   766
  \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   767
  by (simp add: bit_iff_odd_drop_bit)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   768
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   769
lemma div_push_bit_of_1_eq_drop_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   770
  \<open>a div push_bit n 1 = drop_bit n a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   771
  by (simp add: push_bit_eq_mult drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   772
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   773
lemma bits_ident:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   774
  "push_bit n (drop_bit n a) + take_bit n a = a"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   775
  using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   776
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   777
lemma push_bit_push_bit [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   778
  "push_bit m (push_bit n a) = push_bit (m + n) a"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   779
  by (simp add: push_bit_eq_mult power_add ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   780
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   781
lemma push_bit_0_id [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   782
  "push_bit 0 = id"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   783
  by (simp add: fun_eq_iff push_bit_eq_mult)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   784
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   785
lemma push_bit_of_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   786
  "push_bit n 0 = 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   787
  by (simp add: push_bit_eq_mult)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   788
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   789
lemma push_bit_of_1:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   790
  "push_bit n 1 = 2 ^ n"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   791
  by (simp add: push_bit_eq_mult)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   792
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   793
lemma push_bit_Suc [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   794
  "push_bit (Suc n) a = push_bit n (a * 2)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   795
  by (simp add: push_bit_eq_mult ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   796
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   797
lemma push_bit_double:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   798
  "push_bit n (a * 2) = push_bit n a * 2"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   799
  by (simp add: push_bit_eq_mult ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   800
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   801
lemma push_bit_add:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   802
  "push_bit n (a + b) = push_bit n a + push_bit n b"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   803
  by (simp add: push_bit_eq_mult algebra_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   804
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   805
lemma push_bit_numeral [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   806
  \<open>push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   807
  by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   808
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   809
lemma take_bit_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   810
  "take_bit 0 a = 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   811
  by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   812
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   813
lemma take_bit_Suc:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   814
  \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   815
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   816
  have \<open>take_bit (Suc n) (a div 2 * 2 + of_bool (odd a)) = take_bit n (a div 2) * 2 + of_bool (odd a)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   817
    using even_succ_mod_exp [of \<open>2 * (a div 2)\<close> \<open>Suc n\<close>]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   818
      mult_exp_mod_exp_eq [of 1 \<open>Suc n\<close> \<open>a div 2\<close>]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   819
    by (auto simp add: take_bit_eq_mod ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   820
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   821
    using div_mult_mod_eq [of a 2] by (simp add: mod_2_eq_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   822
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   823
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   824
lemma take_bit_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   825
  \<open>take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   826
  by (cases n) (simp_all add: take_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   827
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   828
lemma take_bit_Suc_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   829
  \<open>take_bit (Suc 0) a = a mod 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   830
  by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   831
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   832
lemma take_bit_of_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   833
  "take_bit n 0 = 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   834
  by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   835
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   836
lemma take_bit_of_1 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   837
  "take_bit n 1 = of_bool (n > 0)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   838
  by (cases n) (simp_all add: take_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   839
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   840
lemma drop_bit_of_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   841
  "drop_bit n 0 = 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   842
  by (simp add: drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   843
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   844
lemma drop_bit_of_1 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   845
  "drop_bit n 1 = of_bool (n = 0)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   846
  by (simp add: drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   847
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   848
lemma drop_bit_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   849
  "drop_bit 0 = id"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   850
  by (simp add: fun_eq_iff drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   851
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   852
lemma drop_bit_Suc:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   853
  "drop_bit (Suc n) a = drop_bit n (a div 2)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   854
  using div_exp_eq [of a 1] by (simp add: drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   855
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   856
lemma drop_bit_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   857
  "drop_bit n a = (if n = 0 then a else drop_bit (n - 1) (a div 2))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   858
  by (cases n) (simp_all add: drop_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   859
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   860
lemma drop_bit_half:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   861
  "drop_bit n (a div 2) = drop_bit n a div 2"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   862
  by (induction n arbitrary: a) (simp_all add: drop_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   863
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   864
lemma drop_bit_of_bool [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   865
  "drop_bit n (of_bool b) = of_bool (n = 0 \<and> b)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   866
  by (cases n) simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   867
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   868
lemma even_take_bit_eq [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   869
  \<open>even (take_bit n a) \<longleftrightarrow> n = 0 \<or> even a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   870
  by (simp add: take_bit_rec [of n a])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   871
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   872
lemma take_bit_take_bit [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   873
  "take_bit m (take_bit n a) = take_bit (min m n) a"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   874
  by (simp add: take_bit_eq_mod mod_exp_eq ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   875
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   876
lemma drop_bit_drop_bit [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   877
  "drop_bit m (drop_bit n a) = drop_bit (m + n) a"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   878
  by (simp add: drop_bit_eq_div power_add div_exp_eq ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   879
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   880
lemma push_bit_take_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   881
  "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   882
  apply (simp add: push_bit_eq_mult take_bit_eq_mod power_add ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   883
  using mult_exp_mod_exp_eq [of m \<open>m + n\<close> a] apply (simp add: ac_simps power_add)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   884
  done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   885
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   886
lemma take_bit_push_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   887
  "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   888
proof (cases "m \<le> n")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   889
  case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   890
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   891
    apply (simp add:)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   892
    apply (simp_all add: push_bit_eq_mult take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   893
    apply (auto dest!: le_Suc_ex simp add: power_add ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   894
    using mult_exp_mod_exp_eq [of m m \<open>a * 2 ^ n\<close> for n]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   895
    apply (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   896
    done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   897
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   898
  case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   899
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   900
    using push_bit_take_bit [of n "m - n" a]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   901
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   902
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   903
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   904
lemma take_bit_drop_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   905
  "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   906
  by (simp add: drop_bit_eq_div take_bit_eq_mod ac_simps div_exp_mod_exp_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   907
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   908
lemma drop_bit_take_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   909
  "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   910
proof (cases "m \<le> n")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   911
  case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   912
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   913
    using take_bit_drop_bit [of "n - m" m a] by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   914
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   915
  case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   916
  then obtain q where \<open>m = n + q\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   917
    by (auto simp add: not_le dest: less_imp_Suc_add)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   918
  then have \<open>drop_bit m (take_bit n a) = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   919
    using div_exp_eq [of \<open>a mod 2 ^ n\<close> n q]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   920
    by (simp add: take_bit_eq_mod drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   921
  with False show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   922
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   923
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   924
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   925
lemma even_push_bit_iff [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   926
  \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   927
  by (simp add: push_bit_eq_mult) auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   928
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   929
lemma bit_push_bit_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   930
  \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> 2 ^ n \<noteq> 0 \<and> bit a (n - m)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   931
  by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   932
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   933
lemma bit_drop_bit_eq [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   934
  \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   935
  by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   936
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   937
lemma bit_take_bit_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   938
  \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   939
  by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   940
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   941
lemma stable_imp_drop_bit_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   942
  \<open>drop_bit n a = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   943
  if \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   944
  by (induction n) (simp_all add: that drop_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   945
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   946
lemma stable_imp_take_bit_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   947
  \<open>take_bit n a = (if even a then 0 else 2 ^ n - 1)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   948
    if \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   949
proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   950
  fix m
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   951
  assume \<open>2 ^ m \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   952
  with that show \<open>bit (take_bit n a) m \<longleftrightarrow> bit (if even a then 0 else 2 ^ n - 1) m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   953
    by (simp add: bit_take_bit_iff bit_mask_iff stable_imp_bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   954
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   955
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   956
lemma exp_dvdE:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   957
  assumes \<open>2 ^ n dvd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   958
  obtains b where \<open>a = push_bit n b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   959
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   960
  from assms obtain b where \<open>a = 2 ^ n * b\<close> ..
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   961
  then have \<open>a = push_bit n b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   962
    by (simp add: push_bit_eq_mult ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   963
  with that show thesis .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   964
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   965
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   966
lemma take_bit_eq_0_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   967
  \<open>take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   968
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   969
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   970
  then show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   971
    by (simp add: take_bit_eq_mod mod_0_imp_dvd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   972
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   973
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   974
  then obtain b where \<open>a = push_bit n b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   975
    by (rule exp_dvdE)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   976
  then show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   977
    by (simp add: take_bit_push_bit)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   978
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   979
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   980
lemma take_bit_tightened:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   981
  \<open>take_bit m a = take_bit m b\<close> if \<open>take_bit n a = take_bit n b\<close> and \<open>m \<le> n\<close> 
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   982
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   983
  from that have \<open>take_bit m (take_bit n a) = take_bit m (take_bit n b)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   984
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   985
  then have \<open>take_bit (min m n) a = take_bit (min m n) b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   986
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   987
  with that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   988
    by (simp add: min_def)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   989
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   990
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   991
lemma take_bit_eq_self_iff_drop_bit_eq_0:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   992
  \<open>take_bit n a = a \<longleftrightarrow> drop_bit n a = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   993
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   994
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   995
  show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   996
  proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   997
    fix m
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   998
    from \<open>?P\<close> have \<open>a = take_bit n a\<close> ..
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   999
    also have \<open>\<not> bit (take_bit n a) (n + m)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1000
      unfolding bit_simps
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1001
      by (simp add: bit_simps) 
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1002
    finally show \<open>bit (drop_bit n a) m \<longleftrightarrow> bit 0 m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1003
      by (simp add: bit_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1004
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1005
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1006
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1007
  show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1008
  proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1009
    fix m
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1010
    from \<open>?Q\<close> have \<open>\<not> bit (drop_bit n a) (m - n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1011
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1012
    then have \<open> \<not> bit a (n + (m - n))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1013
      by (simp add: bit_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1014
    then show \<open>bit (take_bit n a) m \<longleftrightarrow> bit a m\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1015
      by (cases \<open>m < n\<close>) (auto simp add: bit_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1016
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1017
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1018
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1019
lemma drop_bit_exp_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1020
  \<open>drop_bit m (2 ^ n) = of_bool (m \<le> n \<and> 2 ^ n \<noteq> 0) * 2 ^ (n - m)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1021
  by (rule bit_eqI) (auto simp add: bit_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1022
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1023
lemma take_bit_and [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1024
  \<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
  1025
  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
  1026
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1027
lemma take_bit_or [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1028
  \<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
  1029
  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
  1030
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1031
lemma take_bit_xor [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1032
  \<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
  1033
  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
  1034
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1035
lemma push_bit_and [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1036
  \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1037
  by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1038
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1039
lemma push_bit_or [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1040
  \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1041
  by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1042
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1043
lemma push_bit_xor [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1044
  \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1045
  by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1046
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1047
lemma drop_bit_and [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1048
  \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1049
  by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1050
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1051
lemma drop_bit_or [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1052
  \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1053
  by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1054
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1055
lemma drop_bit_xor [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1056
  \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1057
  by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1058
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1059
lemma bit_mask_iff [bit_simps]:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1060
  \<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
  1061
  by (simp add: mask_eq_exp_minus_1 bit_mask_iff)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1062
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1063
lemma even_mask_iff:
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1064
  \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1065
  using bit_mask_iff [of n 0] by auto
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1066
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1067
lemma mask_0 [simp]:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1068
  \<open>mask 0 = 0\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1069
  by (simp add: mask_eq_exp_minus_1)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1070
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1071
lemma mask_Suc_0 [simp]:
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1072
  \<open>mask (Suc 0) = 1\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1073
  by (simp add: mask_eq_exp_minus_1 add_implies_diff sym)
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1074
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1075
lemma mask_Suc_exp:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1076
  \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1077
  by (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1078
    (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
  1079
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1080
lemma mask_Suc_double:
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1081
  \<open>mask (Suc n) = 1 OR 2 * mask n\<close>
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1082
proof (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1083
  fix q
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1084
  assume \<open>2 ^ q \<noteq> 0\<close>
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1085
  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
  1086
    by (cases q)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1087
      (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
  1088
qed
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1089
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1090
lemma mask_numeral:
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1091
  \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1092
  by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps)
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1093
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1094
lemma take_bit_mask [simp]:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1095
  \<open>take_bit m (mask n) = mask (min m n)\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1096
  by (rule bit_eqI) (simp add: bit_simps)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1097
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
  1098
lemma take_bit_eq_mask:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1099
  \<open>take_bit n a = a AND mask n\<close>
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1100
  by (rule bit_eqI)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1101
    (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff)
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1102
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1103
lemma or_eq_0_iff:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1104
  \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72611
diff changeset
  1105
  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
  1106
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1107
lemma disjunctive_add:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1108
  \<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
  1109
  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
  1110
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1111
lemma bit_iff_and_drop_bit_eq_1:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1112
  \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1113
  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
  1114
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1115
lemma bit_iff_and_push_bit_not_eq_0:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1116
  \<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
  1117
  apply (cases \<open>2 ^ n = 0\<close>)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1118
  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
  1119
  apply (simp_all add: bit_exp_iff)
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1120
  done
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1121
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1122
lemmas set_bit_def = set_bit_eq_or
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1123
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1124
lemma bit_set_bit_iff [bit_simps]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1125
  \<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
  1126
  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
  1127
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1128
lemma even_set_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1129
  \<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
  1130
  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
  1131
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1132
lemma even_unset_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1133
  \<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
  1134
  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
  1135
73789
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1136
lemma and_exp_eq_0_iff_not_bit:
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1137
  \<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
  1138
proof
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1139
  assume ?Q
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1140
  then show ?P
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1141
    by (auto intro: bit_eqI simp add: bit_simps)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1142
next
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1143
  assume ?P
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1144
  show ?Q
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1145
  proof (rule notI)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1146
    assume \<open>bit a n\<close>
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1147
    then have \<open>a AND 2 ^ n = 2 ^ n\<close>
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1148
      by (auto intro: bit_eqI simp add: bit_simps)
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1149
    with \<open>?P\<close> show False
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1150
      using \<open>bit a n\<close> exp_eq_0_imp_not_bit by auto
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1151
  qed
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1152
qed
aab7975fa070 more lemmas
haftmann
parents: 73682
diff changeset
  1153
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1154
lemmas flip_bit_def = flip_bit_eq_xor
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1155
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1156
lemma bit_flip_bit_iff [bit_simps]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1157
  \<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
  1158
  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
  1159
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1160
lemma even_flip_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1161
  \<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
  1162
  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
  1163
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1164
lemma set_bit_0 [simp]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1165
  \<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
  1166
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1167
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1168
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1169
  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
  1170
    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
  1171
      (cases m, simp_all add: bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1172
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1173
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1174
lemma set_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1175
  \<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
  1176
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1177
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1178
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1179
  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
  1180
  proof (cases m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1181
    case 0
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1182
    then show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1183
      by (simp add: even_set_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1184
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1185
    case (Suc m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1186
    with * have \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1187
      using mult_2 by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1188
    show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1189
      by (cases a rule: parity_cases)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1190
        (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
  1191
        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
  1192
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1193
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1194
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1195
lemma unset_bit_0 [simp]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1196
  \<open>unset_bit 0 a = 2 * (a div 2)\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1197
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1198
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1199
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1200
  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
  1201
    by (simp add: bit_unset_bit_iff bit_double_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1202
      (cases m, simp_all add: bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1203
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1204
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1205
lemma unset_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1206
  \<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
  1207
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1208
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1209
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1210
  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
  1211
  proof (cases m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1212
    case 0
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1213
    then show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1214
      by (simp add: even_unset_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1215
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1216
    case (Suc m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1217
    show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1218
      by (cases a rule: parity_cases)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1219
        (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
  1220
         simp_all add: Suc bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1221
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1222
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1223
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1224
lemma flip_bit_0 [simp]:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1225
  \<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
  1226
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1227
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1228
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1229
  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
  1230
    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
  1231
      (cases m, simp_all add: bit_Suc)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1232
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1233
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1234
lemma flip_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1235
  \<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
  1236
proof (rule bit_eqI)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1237
  fix m
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1238
  assume *: \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1239
  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
  1240
  proof (cases m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1241
    case 0
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1242
    then show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1243
      by (simp add: even_flip_bit_iff)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1244
  next
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1245
    case (Suc m)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1246
    with * have \<open>2 ^ m \<noteq> 0\<close>
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1247
      using mult_2 by auto
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1248
    show ?thesis
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1249
      by (cases a rule: parity_cases)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1250
        (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
  1251
        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
  1252
  qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1253
qed
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1254
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1255
lemma flip_bit_eq_if:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1256
  \<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
  1257
  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
  1258
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1259
lemma take_bit_set_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1260
  \<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
  1261
  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
  1262
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1263
lemma take_bit_unset_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1264
  \<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
  1265
  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
  1266
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1267
lemma take_bit_flip_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1268
  \<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
  1269
  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
  1270
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1271
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1272
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1273
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
  1274
  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1275
  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
  1276
  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
  1277
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1278
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1279
text \<open>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1280
  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
  1281
  definitional class operation.  Note that \<^const>\<open>not\<close> has no
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1282
  sensible definition for unlimited but only positive bit strings
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1283
  (type \<^typ>\<open>nat\<close>).
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1284
\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1285
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1286
lemma bits_minus_1_mod_2_eq [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1287
  \<open>(- 1) mod 2 = 1\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1288
  by (simp add: mod_2_eq_odd)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1289
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1290
lemma not_eq_complement:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1291
  \<open>NOT a = - a - 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1292
  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
  1293
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1294
lemma minus_eq_not_plus_1:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1295
  \<open>- a = NOT a + 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1296
  using not_eq_complement [of a] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1297
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1298
lemma bit_minus_iff [bit_simps]:
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1299
  \<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
  1300
  by (simp add: minus_eq_not_minus_1 bit_not_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1301
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1302
lemma even_not_iff [simp]:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1303
  \<open>even (NOT a) \<longleftrightarrow> odd a\<close>
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1304
  using bit_not_iff [of a 0] by auto
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1305
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1306
lemma bit_not_exp_iff [bit_simps]:
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1307
  \<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
  1308
  by (auto simp add: bit_not_iff bit_exp_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1309
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1310
lemma bit_minus_1_iff [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1311
  \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1312
  by (simp add: bit_minus_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1313
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1314
lemma bit_minus_exp_iff [bit_simps]:
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1315
  \<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
  1316
  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
  1317
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1318
lemma bit_minus_2_iff [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1319
  \<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
  1320
  by (simp add: bit_minus_iff bit_1_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1321
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1322
lemma not_one [simp]:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1323
  \<open>NOT 1 = - 2\<close>
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1324
  by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1325
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1326
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1327
  by standard (rule bit_eqI, simp add: bit_and_iff)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1328
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1329
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
  1330
  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
  1331
proof -
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1332
  interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1333
    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
  1334
  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
  1335
    by standard
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
  1336
  show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1337
    by (rule ext, rule ext, rule bit_eqI)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1338
      (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
  1339
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1340
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1341
lemma and_eq_not_not_or:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1342
  \<open>a AND b = NOT (NOT a OR NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1343
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1344
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1345
lemma or_eq_not_not_and:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1346
  \<open>a OR b = NOT (NOT a AND NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1347
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1348
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1349
lemma not_add_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1350
  \<open>NOT (a + b) = NOT a - b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1351
  by (simp add: not_eq_complement algebra_simps)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1352
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1353
lemma not_diff_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1354
  \<open>NOT (a - b) = NOT a + b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1355
  using not_add_distrib [of a \<open>- b\<close>] by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1356
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1357
lemma and_eq_minus_1_iff:
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1358
  \<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
  1359
proof
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1360
  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
  1361
  then show \<open>a AND b = - 1\<close>
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72611
diff changeset
  1362
    by simp
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1363
next
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1364
  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
  1365
  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
  1366
  proof -
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1367
    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
  1368
    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
  1369
      by (simp add: bit_eq_iff)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1370
    then show \<open>bit a n\<close> \<open>bit b n\<close>
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72611
diff changeset
  1371
      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
  1372
  qed
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1373
  have \<open>a = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1374
    by (rule bit_eqI) (simp add: *)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1375
  moreover have \<open>b = - 1\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1376
    by (rule bit_eqI) (simp add: *)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1377
  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
  1378
    by simp
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1379
qed
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1380
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1381
lemma disjunctive_diff:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1382
  \<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
  1383
proof -
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1384
  have \<open>NOT a + b = NOT a OR b\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1385
    by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1386
  then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1387
    by simp
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1388
  then show ?thesis
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1389
    by (simp add: not_add_distrib)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1390
qed
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1391
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1392
lemma push_bit_minus:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1393
  \<open>push_bit n (- a) = - push_bit n a\<close>
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1394
  by (simp add: push_bit_eq_mult)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1395
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1396
lemma take_bit_not_take_bit:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1397
  \<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
  1398
  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
  1399
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1400
lemma take_bit_not_iff:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1401
  \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1402
  apply (simp add: bit_eq_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1403
  apply (simp add: bit_not_iff bit_take_bit_iff bit_exp_iff)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1404
  apply (use exp_eq_0_imp_not_bit in blast)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1405
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1406
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1407
lemma take_bit_not_eq_mask_diff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1408
  \<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
  1409
proof -
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1410
  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
  1411
    by (simp add: take_bit_not_take_bit)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1412
  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
  1413
    by (simp add: take_bit_eq_mask ac_simps)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1414
  also have \<open>\<dots> = mask n - take_bit n a\<close>
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1415
    by (subst disjunctive_diff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1416
      (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
  1417
  finally show ?thesis
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1418
    by simp
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1419
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1420
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1421
lemma mask_eq_take_bit_minus_one:
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1422
  \<open>mask n = take_bit n (- 1)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1423
  by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1424
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1425
lemma take_bit_minus_one_eq_mask:
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1426
  \<open>take_bit n (- 1) = mask n\<close>
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1427
  by (simp add: mask_eq_take_bit_minus_one)
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1428
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1429
lemma minus_exp_eq_not_mask:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1430
  \<open>- (2 ^ n) = NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1431
  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
  1432
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1433
lemma push_bit_minus_one_eq_not_mask:
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1434
  \<open>push_bit n (- 1) = NOT (mask n)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1435
  by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1436
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1437
lemma take_bit_not_mask_eq_0:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1438
  \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1439
  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
  1440
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1441
lemma unset_bit_eq_and_not:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1442
  \<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
  1443
  by (rule bit_eqI) (auto simp add: bit_simps)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
  1444
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1445
lemmas unset_bit_def = unset_bit_eq_and_not
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1446
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1447
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1448
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1449
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1450
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
  1451
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1452
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
  1453
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1454
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1455
definition not_int :: \<open>int \<Rightarrow> int\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1456
  where \<open>not_int k = - k - 1\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1457
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1458
lemma not_int_rec:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1459
  \<open>NOT k = of_bool (even k) + 2 * NOT (k div 2)\<close> for k :: int
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1460
  by (auto simp add: not_int_def elim: oddE)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1461
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1462
lemma even_not_iff_int:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1463
  \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1464
  by (simp add: not_int_def)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1465
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1466
lemma not_int_div_2:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
  1467
  \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1468
  by (cases k) (simp_all add: not_int_def divide_int_def nat_add_distrib)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1469
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1470
lemma bit_not_int_iff [bit_simps]:
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1471
  \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1472
  for k :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1473
  by (simp add: bit_not_int_iff' not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1474
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1475
function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1476
  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
  1477
    then - of_bool (odd k \<and> odd l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1478
    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
  1479
  by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1480
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1481
termination proof (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1482
  show \<open>wf (measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>)))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1483
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1484
  show \<open>((k div 2, l div 2), k, l) \<in> measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1485
    if \<open>\<not> (k \<in> {0, - 1} \<and> l \<in> {0, - 1})\<close> for k l
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1486
  proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1487
    have less_eq: \<open>\<bar>k div 2\<bar> \<le> \<bar>k\<bar>\<close> for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1488
      by (cases k) (simp_all add: divide_int_def nat_add_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1489
    have less: \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close> if \<open>k \<notin> {0, - 1}\<close> for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1490
    proof (cases k)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1491
      case (nonneg n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1492
      with that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1493
        by (simp add: int_div_less_self)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1494
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1495
      case (neg n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1496
      with that have \<open>n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1497
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1498
      then have \<open>n div 2 < n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1499
        by (simp add: div_less_iff_less_mult)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1500
      with neg that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1501
        by (simp add: divide_int_def nat_add_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1502
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1503
    from that have *: \<open>k \<notin> {0, - 1} \<or> l \<notin> {0, - 1}\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1504
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1505
    then have \<open>0 < \<bar>k\<bar> + \<bar>l\<bar>\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1506
      by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1507
    moreover from * have \<open>\<bar>k div 2\<bar> + \<bar>l div 2\<bar> < \<bar>k\<bar> + \<bar>l\<bar>\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1508
    proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1509
      assume \<open>k \<notin> {0, - 1}\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1510
      then have \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1511
        by (rule less)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1512
      with less_eq [of l] show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1513
        by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1514
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1515
      assume \<open>l \<notin> {0, - 1}\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1516
      then have \<open>\<bar>l div 2\<bar> < \<bar>l\<bar>\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1517
        by (rule less)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1518
      with less_eq [of k] show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1519
        by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1520
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1521
    ultimately show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1522
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1523
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1524
qed
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1525
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1526
declare and_int.simps [simp del]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1527
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1528
lemma and_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1529
  \<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
  1530
    for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1531
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
  1532
  case True
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1533
  then show ?thesis
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1534
    by auto (simp_all add: and_int.simps)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1535
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1536
  case False
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1537
  then show ?thesis
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1538
    by (auto simp add: ac_simps and_int.simps [of k l])
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1539
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1540
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1541
lemma bit_and_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1542
  \<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
  1543
proof (induction n arbitrary: k l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1544
  case 0
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1545
  then show ?case
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1546
    by (simp add: and_int_rec [of k l])
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1547
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1548
  case (Suc n)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1549
  then show ?case
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1550
    by (simp add: and_int_rec [of k l] bit_Suc)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1551
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1552
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1553
lemma even_and_iff_int:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1554
  \<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
  1555
  using bit_and_int_iff [of k l 0] by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1556
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1557
definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1558
  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
  1559
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1560
lemma or_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1561
  \<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
  1562
  for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1563
  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
  1564
  by (simp add: or_int_def even_not_iff_int not_int_div_2)
73535
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1565
    (simp_all add: not_int_def)
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1566
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1567
lemma bit_or_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1568
  \<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
  1569
  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
  1570
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1571
definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1572
  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
  1573
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1574
lemma xor_int_rec:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1575
  \<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
  1576
  for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1577
  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
  1578
    (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
  1579
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1580
lemma bit_xor_int_iff:
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1581
  \<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
  1582
  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
  1583
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1584
definition mask_int :: \<open>nat \<Rightarrow> int\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1585
  where \<open>mask n = (2 :: int) ^ n - 1\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1586
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1587
definition push_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1588
  where \<open>push_bit_int n k = k * 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1589
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1590
definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1591
  where \<open>drop_bit_int n k = k div 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1592
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1593
definition take_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1594
  where \<open>take_bit_int n k = k mod 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1595
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1596
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
  1597
  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
  1598
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1599
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
  1600
  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
  1601
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1602
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
  1603
  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
  1604
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1605
instance proof
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1606
  fix k l :: int and m n :: nat
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1607
  show \<open>- k = NOT (k - 1)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1608
    by (simp add: not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1609
  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
  1610
    by (fact bit_and_int_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1611
  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
  1612
    by (fact bit_or_int_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1613
  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
  1614
    by (fact bit_xor_int_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1615
  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
  1616
  proof -
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1617
    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
  1618
      by (simp add: unset_bit_int_def)
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1619
    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
  1620
      by (simp add: not_int_def)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1621
    finally show ?thesis by (simp only: bit_simps bit_and_int_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1622
      (auto simp add: bit_simps bit_not_int_iff' push_bit_int_def)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1623
  qed
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1624
qed (simp_all add: bit_not_int_iff mask_int_def set_bit_int_def flip_bit_int_def
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1625
  push_bit_int_def drop_bit_int_def take_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
  1626
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1627
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1628
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1629
lemma bit_push_bit_iff_int:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1630
  \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1631
  by (auto simp add: bit_push_bit_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1632
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1633
lemma take_bit_nonnegative [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1634
  \<open>take_bit n k \<ge> 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1635
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1636
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1637
lemma not_take_bit_negative [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1638
  \<open>\<not> take_bit n k < 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1639
  by (simp add: not_less)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1640
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1641
lemma take_bit_int_less_exp [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1642
  \<open>take_bit n k < 2 ^ n\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1643
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1644
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1645
lemma take_bit_int_eq_self_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1646
  \<open>take_bit n k = k \<longleftrightarrow> 0 \<le> k \<and> k < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1647
  for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1648
proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1649
  assume ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1650
  moreover note take_bit_int_less_exp [of n k] take_bit_nonnegative [of n k]
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1651
  ultimately show ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1652
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1653
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1654
  assume ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1655
  then show ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1656
    by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1657
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1658
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1659
lemma take_bit_int_eq_self:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1660
  \<open>take_bit n k = k\<close> if \<open>0 \<le> k\<close> \<open>k < 2 ^ n\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1661
  using that by (simp add: take_bit_int_eq_self_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1662
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1663
lemma mask_half_int:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1664
  \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1665
  by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  1666
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1667
lemma mask_nonnegative_int [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1668
  \<open>mask n \<ge> (0::int)\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1669
  by (simp add: mask_eq_exp_minus_1)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1670
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1671
lemma not_mask_negative_int [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1672
  \<open>\<not> mask n < (0::int)\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1673
  by (simp add: not_less)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1674
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1675
lemma not_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1676
  \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1677
  by (simp add: not_int_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1678
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1679
lemma not_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1680
  \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1681
  by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1682
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1683
lemma and_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1684
  \<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
  1685
proof (induction k arbitrary: l rule: int_bit_induct)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1686
  case zero
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1687
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1688
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1689
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1690
  case minus
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1691
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1692
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1693
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1694
  case (even k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1695
  then show ?case
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1696
    using and_int_rec [of \<open>k * 2\<close> l]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1697
    by (simp add: pos_imp_zdiv_nonneg_iff zero_le_mult_iff)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1698
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1699
  case (odd k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1700
  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
  1701
    by simp
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1702
  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>
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1703
    by simp
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1704
  with and_int_rec [of \<open>1 + k * 2\<close> l]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1705
  show ?case
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1706
    by (auto simp add: zero_le_mult_iff not_le)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1707
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1708
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1709
lemma and_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1710
  \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1711
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1712
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1713
lemma and_less_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1714
  \<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
  1715
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1716
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1717
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1718
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1719
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1720
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1721
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1722
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1723
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1724
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1725
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1726
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1727
    by (simp add: and_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1728
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1729
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1730
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1731
  show ?case
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1732
    by (simp add: and_int_rec [of _ l]) linarith
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1733
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1734
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1735
lemma or_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1736
  \<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
  1737
  by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1738
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1739
lemma or_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1740
  \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1741
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1742
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1743
lemma or_greater_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1744
  \<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
  1745
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1746
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1747
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1748
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1749
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1750
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1751
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1752
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1753
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1754
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1755
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1756
  show ?case
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1757
    by (simp add: or_int_rec [of _ l]) linarith
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1758
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1759
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1760
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1761
  show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1762
    by (simp add: or_int_rec [of _ l])
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1763
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1764
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1765
lemma xor_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1766
  \<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
  1767
  by (simp only: bit.xor_def or_nonnegative_int_iff) auto
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1768
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1769
lemma xor_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1770
  \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1771
  by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1772
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1773
lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1774
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1775
  assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1776
  shows \<open>x OR y < 2 ^ n\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1777
using assms proof (induction x arbitrary: y n rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1778
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1779
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1780
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1781
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1782
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1783
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1784
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1785
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1786
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1787
  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
  1788
  show ?case 
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1789
    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
  1790
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1791
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1792
  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
  1793
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1794
    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
  1795
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1796
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1797
lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1798
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1799
  assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1800
  shows \<open>x XOR y < 2 ^ n\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1801
using assms proof (induction x arbitrary: y n rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1802
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1803
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1804
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1805
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1806
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1807
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1808
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1809
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1810
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1811
  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
  1812
  show ?case 
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1813
    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
  1814
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1815
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1816
  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
  1817
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1818
    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
  1819
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1820
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1821
lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1822
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1823
  assumes \<open>0 \<le> x\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1824
  shows \<open>0 \<le> x AND y\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1825
  using assms by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1826
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1827
lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1828
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1829
  assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1830
  shows \<open>0 \<le> x OR y\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1831
  using assms by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1832
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1833
lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1834
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1835
  assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1836
  shows \<open>0 \<le> x XOR y\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1837
  using assms by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1838
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1839
lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1840
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1841
  assumes \<open>0 \<le> x\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1842
  shows \<open>x AND y \<le> x\<close>
73535
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1843
using assms proof (induction x arbitrary: y rule: int_bit_induct)
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1844
  case (odd k)
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1845
  then have \<open>k AND y div 2 \<le> k\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1846
    by simp
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1847
  then show ?case 
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1848
    by (simp add: and_int_rec [of \<open>1 + _ * 2\<close>])
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  1849
qed (simp_all add: and_int_rec [of \<open>_ * 2\<close>])
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1850
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1851
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
  1852
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
  1853
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1854
lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1855
  fixes x y :: int
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1856
  assumes \<open>0 \<le> y\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1857
  shows \<open>x AND y \<le> y\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1858
  using assms AND_upper1 [of y x] by (simp add: ac_simps)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1859
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1860
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
  1861
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
  1862
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1863
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
  1864
proof (induction x arbitrary: y rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1865
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1866
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1867
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1868
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1869
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1870
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1871
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1872
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1873
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1874
  from even.IH [of \<open>y div 2\<close>]
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1875
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1876
    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
  1877
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1878
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1879
  from odd.IH [of \<open>y div 2\<close>]
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1880
  show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1881
    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
  1882
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  1883
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1884
lemma push_bit_minus_one:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1885
  "push_bit n (- 1 :: int) = - (2 ^ n)"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1886
  by (simp add: push_bit_eq_mult)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1887
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1888
lemma minus_1_div_exp_eq_int:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1889
  \<open>- 1 div (2 :: int) ^ n = - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1890
  by (induction n) (use div_exp_eq [symmetric, of \<open>- 1 :: int\<close> 1] in \<open>simp_all add: ac_simps\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1891
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1892
lemma drop_bit_minus_one [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1893
  \<open>drop_bit n (- 1 :: int) = - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1894
  by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1895
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1896
lemma take_bit_Suc_from_most:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1897
  \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1898
  by (simp only: take_bit_eq_mod power_Suc2) (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one zmod_zmult2_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1899
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1900
lemma take_bit_minus:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1901
  \<open>take_bit n (- take_bit n k) = take_bit n (- k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1902
    for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1903
  by (simp add: take_bit_eq_mod mod_minus_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1904
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1905
lemma take_bit_diff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1906
  \<open>take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1907
    for k l :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1908
  by (simp add: take_bit_eq_mod mod_diff_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1909
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1910
lemma bit_imp_take_bit_positive:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1911
  \<open>0 < take_bit m k\<close> if \<open>n < m\<close> and \<open>bit k n\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1912
proof (rule ccontr)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1913
  assume \<open>\<not> 0 < take_bit m k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1914
  then have \<open>take_bit m k = 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1915
    by (auto simp add: not_less intro: order_antisym)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1916
  then have \<open>bit (take_bit m k) n = bit 0 n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1917
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1918
  with that show False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1919
    by (simp add: bit_take_bit_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1920
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1921
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1922
lemma take_bit_mult:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1923
  \<open>take_bit n (take_bit n k * take_bit n l) = take_bit n (k * l)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1924
  for k l :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1925
  by (simp add: take_bit_eq_mod mod_mult_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1926
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1927
lemma (in ring_1) of_nat_nat_take_bit_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1928
  \<open>of_nat (nat (take_bit n k)) = of_int (take_bit n k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1929
  by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1930
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1931
lemma take_bit_minus_small_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1932
  \<open>take_bit n (- k) = 2 ^ n - k\<close> if \<open>0 < k\<close> \<open>k \<le> 2 ^ n\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1933
proof -
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1934
  define m where \<open>m = nat k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1935
  with that have \<open>k = int m\<close> and \<open>0 < m\<close> and \<open>m \<le> 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1936
    by simp_all
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1937
  have \<open>(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1938
    using \<open>0 < m\<close> by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1939
  then have \<open>int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1940
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1941
  then have \<open>(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1942
    using \<open>m \<le> 2 ^ n\<close> by (simp only: of_nat_mod of_nat_diff) simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1943
  with \<open>k = int m\<close> have \<open>(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1944
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1945
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1946
    by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1947
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1948
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1949
lemma drop_bit_push_bit_int:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1950
  \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1951
  by (cases \<open>m \<le> n\<close>) (auto simp add: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] mult.assoc
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1952
    mult.commute [of k] drop_bit_eq_div push_bit_eq_mult not_le power_add dest!: le_Suc_ex less_imp_Suc_add)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1953
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1954
lemma push_bit_nonnegative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1955
  \<open>push_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1956
  by (simp add: push_bit_eq_mult zero_le_mult_iff power_le_zero_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1957
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1958
lemma push_bit_negative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1959
  \<open>push_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1960
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1961
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1962
lemma drop_bit_nonnegative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1963
  \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1964
  by (induction n) (auto simp add: drop_bit_Suc drop_bit_half)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1965
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1966
lemma drop_bit_negative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1967
  \<open>drop_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1968
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1969
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1970
lemma set_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1971
  \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1972
  by (simp add: set_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1973
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1974
lemma set_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1975
  \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1976
  by (simp add: set_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1977
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1978
lemma unset_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1979
  \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1980
  by (simp add: unset_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1981
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1982
lemma unset_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1983
  \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1984
  by (simp add: unset_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1985
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1986
lemma flip_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1987
  \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1988
  by (simp add: flip_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1989
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1990
lemma flip_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1991
  \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1992
  by (simp add: flip_bit_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1993
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1994
lemma set_bit_greater_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1995
  \<open>set_bit n k \<ge> k\<close> for k :: int
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1996
  by (simp add: set_bit_def or_greater_eq)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1997
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1998
lemma unset_bit_less_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  1999
  \<open>unset_bit n k \<le> k\<close> for k :: int
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2000
  by (simp add: unset_bit_def and_less_eq)
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2001
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2002
lemma set_bit_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2003
  \<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
  2004
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2005
  fix m
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2006
  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
  2007
  proof (cases \<open>m = n\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2008
    case True
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2009
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2010
      apply (simp add: bit_set_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2011
      apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2012
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2013
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2014
    case False
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2015
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2016
      apply (clarsimp simp add: bit_set_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2017
      apply (subst disjunctive_add)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2018
      apply (clarsimp simp add: bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2019
      apply (clarsimp simp add: bit_or_iff bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2020
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2021
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2022
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2023
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2024
lemma unset_bit_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2025
  \<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
  2026
proof (rule bit_eqI)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2027
  fix m
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2028
  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
  2029
  proof (cases \<open>m = n\<close>)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2030
    case True
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2031
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2032
      apply (simp add: bit_unset_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2033
      apply (simp add: bit_iff_odd)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2034
      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
  2035
      apply (simp add: dvd_neg_div)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2036
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2037
  next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2038
    case False
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2039
    then show ?thesis
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2040
      apply (clarsimp simp add: bit_unset_bit_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2041
      apply (subst disjunctive_diff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2042
      apply (clarsimp simp add: bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2043
      apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2044
      done
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2045
  qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2046
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2047
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2048
lemma and_int_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2049
  \<open>k AND l = (if k = 0 \<or> l = 0 then 0 else if k = - 1 then l else if l = - 1 then k
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2050
    else (k mod 2) * (l mod 2) + 2 * ((k div 2) AND (l div 2)))\<close> for k l :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2051
  by (auto simp add: and_int_rec [of k l] zmult_eq_1_iff elim: oddE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2052
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2053
lemma or_int_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2054
  \<open>k OR l = (if k = - 1 \<or> l = - 1 then - 1 else if k = 0 then l else if l = 0 then k
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2055
    else max (k mod 2) (l mod 2) + 2 * ((k div 2) OR (l div 2)))\<close> for k l :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2056
  by (auto simp add: or_int_rec [of k l] elim: oddE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2057
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2058
lemma xor_int_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2059
  \<open>k XOR l = (if k = - 1 then NOT l else if l = - 1 then NOT k else if k = 0 then l else if l = 0 then k
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2060
    else \<bar>k mod 2 - l mod 2\<bar> + 2 * ((k div 2) XOR (l div 2)))\<close> for k l :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2061
  by (auto simp add: xor_int_rec [of k l] not_int_def elim!: oddE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2062
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2063
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2064
subsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2065
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2066
instantiation nat :: semiring_bit_operations
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2067
begin
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2068
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2069
definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2070
  where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2071
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2072
definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2073
  where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2074
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2075
definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2076
  where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2077
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2078
definition mask_nat :: \<open>nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2079
  where \<open>mask n = (2 :: nat) ^ n - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2080
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2081
definition push_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2082
  where \<open>push_bit_nat n m = m * 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2083
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2084
definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2085
  where \<open>drop_bit_nat n m = m div 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2086
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2087
definition take_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2088
  where \<open>take_bit_nat n m = m mod 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2089
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2090
definition set_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2091
  where \<open>set_bit m n = n OR push_bit m 1\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2092
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2093
definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2094
  where \<open>unset_bit m n = nat (unset_bit m (int n))\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2095
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2096
definition flip_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2097
  where \<open>flip_bit m n = n XOR push_bit m 1\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2098
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2099
instance proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2100
  fix m n q :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2101
  show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2102
    by (simp add: and_nat_def bit_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2103
  show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2104
    by (simp add: or_nat_def bit_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2105
  show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2106
    by (simp add: xor_nat_def bit_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2107
  show \<open>bit (unset_bit m n) q \<longleftrightarrow> bit n q \<and> m \<noteq> q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2108
    by (simp add: unset_bit_nat_def bit_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2109
qed (simp_all add: mask_nat_def set_bit_nat_def flip_bit_nat_def push_bit_nat_def drop_bit_nat_def take_bit_nat_def)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2110
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2111
end
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2112
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2113
lemma take_bit_nat_less_exp [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2114
  \<open>take_bit n m < 2 ^ n\<close> for n m ::nat 
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2115
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2116
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2117
lemma take_bit_nat_eq_self_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2118
  \<open>take_bit n m = m \<longleftrightarrow> m < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2119
  for n m :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2120
proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2121
  assume ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2122
  moreover note take_bit_nat_less_exp [of n m]
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2123
  ultimately show ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2124
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2125
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2126
  assume ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2127
  then show ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2128
    by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2129
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2130
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2131
lemma take_bit_nat_eq_self:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2132
  \<open>take_bit n m = m\<close> if \<open>m < 2 ^ n\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2133
  using that by (simp add: take_bit_nat_eq_self_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2134
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2135
lemma take_bit_nat_less_eq_self [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2136
  \<open>take_bit n m \<le> m\<close> for n m :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2137
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2138
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2139
lemma take_bit_nat_less_self_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2140
  \<open>take_bit n m < m \<longleftrightarrow> 2 ^ n \<le> m\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2141
  for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2142
proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2143
  assume ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2144
  then have \<open>take_bit n m \<noteq> m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2145
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2146
  then show \<open>?Q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2147
    by (simp add: take_bit_nat_eq_self_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2148
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2149
  have \<open>take_bit n m < 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2150
    by (fact take_bit_nat_less_exp)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2151
  also assume ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2152
  finally show ?P .
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2153
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2154
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2155
lemma bit_push_bit_iff_nat:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2156
  \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2157
  by (auto simp add: bit_push_bit_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2158
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2159
lemma and_nat_rec:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2160
  \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2161
  apply (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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2162
  apply (subst nat_add_distrib)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2163
    apply auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2164
  done
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2165
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2166
lemma or_nat_rec:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2167
  \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2168
  apply (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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2169
  apply (subst nat_add_distrib)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2170
    apply auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2171
  done
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2172
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2173
lemma xor_nat_rec:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2174
  \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2175
  apply (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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2176
  apply (subst nat_add_distrib)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2177
    apply auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2178
  done
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2179
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2180
lemma Suc_0_and_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2181
  \<open>Suc 0 AND n = n mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2182
  using one_and_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2183
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2184
lemma and_Suc_0_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2185
  \<open>n AND Suc 0 = n mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2186
  using and_one_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2187
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2188
lemma Suc_0_or_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2189
  \<open>Suc 0 OR n = n + of_bool (even n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2190
  using one_or_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2191
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2192
lemma or_Suc_0_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2193
  \<open>n OR Suc 0 = n + of_bool (even n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2194
  using or_one_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2195
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2196
lemma Suc_0_xor_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2197
  \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2198
  using one_xor_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2199
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2200
lemma xor_Suc_0_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2201
  \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2202
  using xor_one_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2203
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2204
lemma and_nat_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2205
  \<open>m AND n = (if m = 0 \<or> n = 0 then 0 else (m mod 2) * (n mod 2) + 2 * ((m div 2) AND (n div 2)))\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2206
    for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2207
  by (auto simp add: and_nat_rec [of m n] elim: oddE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2208
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2209
lemma or_nat_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2210
  \<open>m OR n = (if m = 0 then n else if n = 0 then m
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2211
    else max (m mod 2) (n mod 2) + 2 * ((m div 2) OR (n div 2)))\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2212
  by (auto simp add: or_nat_rec [of m n] elim: oddE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2213
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2214
lemma xor_nat_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2215
  \<open>m XOR n = (if m = 0 then n else if n = 0 then m
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2216
    else (m mod 2 + n mod 2) mod 2 + 2 * ((m div 2) XOR (n div 2)))\<close> for m n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2217
  by (auto simp add: xor_nat_rec [of m n] elim!: oddE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2218
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2219
lemma [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2220
  \<open>unset_bit 0 m = 2 * (m div 2)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2221
  \<open>unset_bit (Suc n) m = m mod 2 + 2 * unset_bit n (m div 2)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2222
  by (simp_all add: unset_bit_Suc)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2223
  
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2224
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2225
subsection \<open>Common algebraic structure\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2226
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2227
class unique_euclidean_semiring_with_bit_operations =
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2228
  unique_euclidean_semiring_with_nat + semiring_bit_operations
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2229
begin
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2230
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2231
lemma take_bit_of_exp [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2232
  \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2233
  by (simp add: take_bit_eq_mod exp_mod_exp)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2234
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2235
lemma take_bit_of_2 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2236
  \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2237
  using take_bit_of_exp [of n 1] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2238
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2239
lemma take_bit_of_mask:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2240
  \<open>take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2241
  by (simp add: take_bit_eq_mod mask_mod_exp)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2242
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2243
lemma push_bit_eq_0_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2244
  "push_bit n a = 0 \<longleftrightarrow> a = 0"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2245
  by (simp add: push_bit_eq_mult)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2246
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2247
lemma take_bit_add:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2248
  "take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2249
  by (simp add: take_bit_eq_mod mod_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2250
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2251
lemma take_bit_of_1_eq_0_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2252
  "take_bit n 1 = 0 \<longleftrightarrow> n = 0"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2253
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2254
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2255
lemma take_bit_Suc_1 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2256
  \<open>take_bit (Suc n) 1 = 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2257
  by (simp add: take_bit_Suc)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2258
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2259
lemma take_bit_Suc_bit0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2260
  \<open>take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2261
  by (simp add: take_bit_Suc numeral_Bit0_div_2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2262
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2263
lemma take_bit_Suc_bit1 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2264
  \<open>take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2265
  by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2266
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2267
lemma take_bit_numeral_1 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2268
  \<open>take_bit (numeral l) 1 = 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2269
  by (simp add: take_bit_rec [of \<open>numeral l\<close> 1])
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2270
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2271
lemma take_bit_numeral_bit0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2272
  \<open>take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2273
  by (simp add: take_bit_rec numeral_Bit0_div_2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2274
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2275
lemma take_bit_numeral_bit1 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2276
  \<open>take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2277
  by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2278
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2279
lemma drop_bit_Suc_bit0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2280
  \<open>drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2281
  by (simp add: drop_bit_Suc numeral_Bit0_div_2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2282
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2283
lemma drop_bit_Suc_bit1 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2284
  \<open>drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2285
  by (simp add: drop_bit_Suc numeral_Bit1_div_2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2286
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2287
lemma drop_bit_numeral_bit0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2288
  \<open>drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2289
  by (simp add: drop_bit_rec numeral_Bit0_div_2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2290
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2291
lemma drop_bit_numeral_bit1 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2292
  \<open>drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2293
  by (simp add: drop_bit_rec numeral_Bit1_div_2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2294
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2295
lemma drop_bit_of_nat:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2296
  "drop_bit n (of_nat m) = of_nat (drop_bit n m)"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2297
  by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div [of m "2 ^ n"])
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2298
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2299
lemma bit_of_nat_iff_bit [bit_simps]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2300
  \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2301
proof -
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2302
  have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2303
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2304
  also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2305
    by (simp add: of_nat_div)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2306
  finally show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2307
    by (simp add: bit_iff_odd semiring_bits_class.bit_iff_odd)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2308
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2309
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2310
lemma of_nat_drop_bit:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2311
  \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2312
  by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2313
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2314
lemma bit_push_bit_iff_of_nat_iff [bit_simps]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2315
  \<open>bit (push_bit m (of_nat r)) n \<longleftrightarrow> m \<le> n \<and> bit (of_nat r) (n - m)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2316
  by (auto simp add: bit_push_bit_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2317
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2318
lemma take_bit_sum:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2319
  "take_bit n a = (\<Sum>k = 0..<n. push_bit k (of_bool (bit a k)))"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2320
  for n :: nat
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2321
proof (induction n arbitrary: a)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2322
  case 0
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2323
  then show ?case
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2324
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2325
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2326
  case (Suc n)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2327
  have "(\<Sum>k = 0..<Suc n. push_bit k (of_bool (bit a k))) = 
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2328
    of_bool (odd a) + (\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2329
    by (simp add: sum.atLeast_Suc_lessThan ac_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2330
  also have "(\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2331
    = (\<Sum>k = 0..<n. push_bit k (of_bool (bit (a div 2) k))) * 2"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2332
    by (simp only: sum.atLeast_Suc_lessThan_Suc_shift) (simp add: sum_distrib_right push_bit_double drop_bit_Suc bit_Suc)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2333
  finally show ?case
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2334
    using Suc [of "a div 2"] by (simp add: ac_simps take_bit_Suc mod_2_eq_odd)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2335
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2336
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2337
end
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2338
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2339
instance nat :: unique_euclidean_semiring_with_bit_operations ..
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2340
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2341
instance int :: unique_euclidean_semiring_with_bit_operations ..
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2342
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2343
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2344
subsection \<open>More properties\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2345
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2346
lemma take_bit_eq_mask_iff:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2347
  \<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
  2348
  for k :: int
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2349
proof
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2350
  assume ?P
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2351
  then have \<open>take_bit n (take_bit n k + take_bit n 1) = 0\<close>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2352
    by (simp add: mask_eq_exp_minus_1 take_bit_eq_0_iff)
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2353
  then show ?Q
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2354
    by (simp only: take_bit_add)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2355
next
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2356
  assume ?Q
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2357
  then have \<open>take_bit n (k + 1) - 1 = - 1\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2358
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2359
  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
  2360
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2361
  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
  2362
    by (simp add: take_bit_eq_mod mod_simps)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2363
  ultimately show ?P
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2364
    by (simp add: take_bit_minus_one_eq_mask)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2365
qed
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2366
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2367
lemma take_bit_eq_mask_iff_exp_dvd:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2368
  \<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
  2369
  for k :: int
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  2370
  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
  2371
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2372
context ring_bit_operations
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2373
begin
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2374
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2375
lemma even_of_int_iff:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2376
  \<open>even (of_int k) \<longleftrightarrow> even k\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2377
  by (induction k rule: int_bit_induct) simp_all
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2378
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  2379
lemma bit_of_int_iff [bit_simps]:
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2380
  \<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
  2381
proof (cases \<open>(2::'a) ^ n = 0\<close>)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2382
  case True
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2383
  then show ?thesis
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2384
    by (simp add: exp_eq_0_imp_not_bit)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2385
next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2386
  case False
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2387
  then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2388
  proof (induction k arbitrary: n rule: int_bit_induct)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2389
    case zero
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2390
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2391
      by simp
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2392
  next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2393
    case minus
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2394
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2395
      by simp
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2396
  next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2397
    case (even k)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2398
    then show ?case
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2399
      using bit_double_iff [of \<open>of_int k\<close> n] Bit_Operations.bit_double_iff [of k n]
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2400
      by (cases n) (auto simp add: ac_simps dest: mult_not_zero)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2401
  next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2402
    case (odd k)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2403
    then show ?case
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2404
      using bit_double_iff [of \<open>of_int k\<close> n]
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2405
      by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Bit_Operations.bit_Suc dest: mult_not_zero)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2406
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2407
  with False show ?thesis
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2408
    by simp
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2409
qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2410
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2411
lemma push_bit_of_int:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2412
  \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2413
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2414
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2415
lemma of_int_push_bit:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2416
  \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2417
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2418
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2419
lemma take_bit_of_int:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2420
  \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2421
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2422
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2423
lemma of_int_take_bit:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2424
  \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2425
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2426
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2427
lemma of_int_not_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2428
  \<open>of_int (NOT k) = NOT (of_int k)\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2429
  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
  2430
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2431
lemma of_int_and_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2432
  \<open>of_int (k AND l) = of_int k AND of_int l\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2433
  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
  2434
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2435
lemma of_int_or_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2436
  \<open>of_int (k OR l) = of_int k OR of_int l\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2437
  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
  2438
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2439
lemma of_int_xor_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2440
  \<open>of_int (k XOR l) = of_int k XOR of_int l\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2441
  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
  2442
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2443
lemma of_int_mask_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2444
  \<open>of_int (mask n) = mask n\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2445
  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
  2446
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2447
end
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2448
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2449
lemma take_bit_incr_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2450
  \<open>take_bit n (k + 1) = 1 + take_bit n k\<close> if \<open>take_bit n k \<noteq> 2 ^ n - 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2451
  for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2452
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2453
  from that have \<open>2 ^ n \<noteq> k mod 2 ^ n + 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2454
    by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2455
  moreover have \<open>k mod 2 ^ n < 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2456
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2457
  ultimately have *: \<open>k mod 2 ^ n + 1 < 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2458
    by linarith
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2459
  have \<open>(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2460
    by (simp add: mod_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2461
  also have \<open>\<dots> = k mod 2 ^ n + 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2462
    using * by (simp add: zmod_trivial_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2463
  finally have \<open>(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\<close> .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2464
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2465
    by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2466
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2467
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2468
lemma take_bit_decr_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2469
  \<open>take_bit n (k - 1) = take_bit n k - 1\<close> if \<open>take_bit n k \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2470
  for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2471
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2472
  from that have \<open>k mod 2 ^ n \<noteq> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2473
    by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2474
  moreover have \<open>k mod 2 ^ n \<ge> 0\<close> \<open>k mod 2 ^ n < 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2475
    by simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2476
  ultimately have *: \<open>k mod 2 ^ n > 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2477
    by linarith
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2478
  have \<open>(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2479
    by (simp add: mod_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2480
  also have \<open>\<dots> = k mod 2 ^ n - 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2481
    by (simp add: zmod_trivial_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2482
      (use \<open>k mod 2 ^ n < 2 ^ n\<close> * in linarith)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2483
  finally have \<open>(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\<close> .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2484
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2485
    by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2486
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2487
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2488
lemma take_bit_int_greater_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2489
  \<open>k + 2 ^ n \<le> take_bit n k\<close> if \<open>k < 0\<close> for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2490
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2491
  have \<open>k + 2 ^ n \<le> take_bit n (k + 2 ^ n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2492
  proof (cases \<open>k > - (2 ^ n)\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2493
    case False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2494
    then have \<open>k + 2 ^ n \<le> 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2495
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2496
    also note take_bit_nonnegative
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2497
    finally show ?thesis .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2498
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2499
    case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2500
    with that have \<open>0 \<le> k + 2 ^ n\<close> and \<open>k + 2 ^ n < 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2501
      by simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2502
    then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2503
      by (simp only: take_bit_eq_mod mod_pos_pos_trivial)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2504
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2505
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2506
    by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2507
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2508
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2509
lemma take_bit_int_less_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2510
  \<open>take_bit n k \<le> k - 2 ^ n\<close> if \<open>2 ^ n \<le> k\<close> and \<open>n > 0\<close> for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2511
  using that zmod_le_nonneg_dividend [of \<open>k - 2 ^ n\<close> \<open>2 ^ n\<close>]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2512
  by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2513
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2514
lemma take_bit_int_less_eq_self_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2515
  \<open>take_bit n k \<le> k \<longleftrightarrow> 0 \<le> k\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2516
  for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2517
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2518
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2519
  show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2520
  proof (rule ccontr)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2521
    assume \<open>\<not> 0 \<le> k\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2522
    then have \<open>k < 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2523
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2524
    with \<open>?P\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2525
    have \<open>take_bit n k < 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2526
      by (rule le_less_trans)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2527
    then show False
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2528
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2529
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2530
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2531
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2532
  then show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2533
    by (simp add: take_bit_eq_mod zmod_le_nonneg_dividend)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2534
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2535
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2536
lemma take_bit_int_less_self_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2537
  \<open>take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2538
  for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2539
  by (auto simp add: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2540
    intro: order_trans [of 0 \<open>2 ^ n\<close> k])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2541
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2542
lemma take_bit_int_greater_self_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2543
  \<open>k < take_bit n k \<longleftrightarrow> k < 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2544
  for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2545
  using take_bit_int_less_eq_self_iff [of n k] by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2546
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2547
lemma take_bit_int_greater_eq_self_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2548
  \<open>k \<le> take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2549
  for k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2550
  by (auto simp add: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2551
    dest: sym not_sym intro: less_trans [of k 0 \<open>2 ^ n\<close>])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2552
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2553
lemma minus_numeral_inc_eq:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2554
  \<open>- numeral (Num.inc n) = NOT (numeral n :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2555
  by (simp add: not_int_def sub_inc_One_eq add_One)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2556
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2557
lemma sub_one_eq_not_neg:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2558
  \<open>Num.sub n num.One = NOT (- numeral n :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2559
  by (simp add: not_int_def)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2560
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2561
lemma bit_numeral_int_iff [bit_simps]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2562
  \<open>bit (numeral m :: int) n \<longleftrightarrow> bit (numeral m :: nat) n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2563
  using bit_of_nat_iff_bit [of \<open>numeral m\<close> n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2564
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2565
lemma bit_minus_int_iff [bit_simps]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2566
  \<open>bit (- k) n \<longleftrightarrow> \<not> bit (k - 1) n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2567
  for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2568
  using bit_not_int_iff' [of \<open>k - 1\<close>] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2569
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2570
lemma bit_numeral_int_simps [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2571
  \<open>bit (1 :: int) (numeral n) \<longleftrightarrow> bit (0 :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2572
  \<open>bit (numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (numeral w :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2573
  \<open>bit (numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> bit (numeral w :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2574
  \<open>bit (numeral (Num.BitM w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (- numeral w :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2575
  \<open>bit (- numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (- numeral w :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2576
  \<open>bit (- numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (numeral w :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2577
  \<open>bit (- numeral (Num.BitM w) :: int) (numeral n) \<longleftrightarrow> bit (- (numeral w) :: int) (pred_numeral n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2578
  by (simp_all add: bit_1_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq bit_minus_int_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2579
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2580
lemma bit_numeral_Bit0_Suc_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2581
  \<open>bit (numeral (Num.Bit0 m) :: int) (Suc n) \<longleftrightarrow> bit (numeral m :: int) n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2582
  by (simp add: bit_Suc)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2583
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2584
lemma bit_numeral_Bit1_Suc_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2585
  \<open>bit (numeral (Num.Bit1 m) :: int) (Suc n) \<longleftrightarrow> bit (numeral m :: int) n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2586
  by (simp add: bit_Suc)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2587
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2588
lemma int_not_numerals [simp]:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2589
  \<open>NOT (numeral (Num.Bit0 n) :: int) = - numeral (Num.Bit1 n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2590
  \<open>NOT (numeral (Num.Bit1 n) :: int) = - numeral (Num.inc (num.Bit1 n))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2591
  \<open>NOT (numeral (Num.BitM n) :: int) = - numeral (num.Bit0 n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2592
  \<open>NOT (- numeral (Num.Bit0 n) :: int) = numeral (Num.BitM n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2593
  \<open>NOT (- numeral (Num.Bit1 n) :: int) = numeral (Num.Bit0 n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2594
  by (simp_all add: not_int_def add_One inc_BitM_eq) 
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2595
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2596
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
  2597
  operator). Is there a simpler way to do this?\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2598
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2599
context
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2600
begin
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2601
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2602
private lemma eqI:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2603
  \<open>k = l\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2604
  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
  2605
    and even: \<open>even k \<longleftrightarrow> even l\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2606
  for k l :: int
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2607
proof (rule bit_eqI)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2608
  fix n
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2609
  show \<open>bit k n \<longleftrightarrow> bit l n\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2610
  proof (cases n)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2611
    case 0
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2612
    with even show ?thesis
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2613
      by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2614
  next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2615
    case (Suc n)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2616
    with num [of \<open>num_of_nat (Suc n)\<close>] show ?thesis
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2617
      by (simp only: numeral_num_of_nat)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2618
  qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2619
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2620
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2621
lemma int_and_numerals [simp]:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2622
  \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2623
  \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2624
  \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2625
  \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2626
  \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2627
  \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2628
  \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2629
  \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2630
  \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2631
  \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2632
  \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2633
  \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2634
  \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2635
  \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2636
  \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2637
  \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2638
  \<open>(1::int) AND numeral (Num.Bit0 y) = 0\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2639
  \<open>(1::int) AND numeral (Num.Bit1 y) = 1\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2640
  \<open>(1::int) AND - numeral (Num.Bit0 y) = 0\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2641
  \<open>(1::int) AND - numeral (Num.Bit1 y) = 1\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2642
  \<open>numeral (Num.Bit0 x) AND (1::int) = 0\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2643
  \<open>numeral (Num.Bit1 x) AND (1::int) = 1\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2644
  \<open>- numeral (Num.Bit0 x) AND (1::int) = 0\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2645
  \<open>- numeral (Num.Bit1 x) AND (1::int) = 1\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2646
  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
  2647
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2648
lemma int_or_numerals [simp]:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2649
  \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2650
  \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2651
  \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2652
  \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2653
  \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2654
  \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2655
  \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2656
  \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2657
  \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2658
  \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2659
  \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2660
  \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2661
  \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2662
  \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2663
  \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2664
  \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2665
  \<open>(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2666
  \<open>(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2667
  \<open>(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2668
  \<open>(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2669
  \<open>numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2670
  \<open>numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2671
  \<open>- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2672
  \<open>- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2673
  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
  2674
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2675
lemma int_xor_numerals [simp]:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2676
  \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2677
  \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2678
  \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2679
  \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2680
  \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2681
  \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2682
  \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2683
  \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2684
  \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2685
  \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2686
  \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2687
  \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2688
  \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2689
  \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2690
  \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2691
  \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2692
  \<open>(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2693
  \<open>(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2694
  \<open>(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2695
  \<open>(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2696
  \<open>numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2697
  \<open>numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2698
  \<open>- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  2699
  \<open>- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))\<close>
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2700
  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
  2701
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2702
end
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2703
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2704
context semiring_bit_operations
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2705
begin
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2706
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2707
lemma push_bit_of_nat:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2708
  \<open>push_bit n (of_nat m) = of_nat (push_bit n m)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2709
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2710
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2711
lemma of_nat_push_bit:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2712
  \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2713
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2714
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2715
lemma take_bit_of_nat:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2716
  \<open>take_bit n (of_nat m) = of_nat (take_bit n m)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2717
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2718
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2719
lemma of_nat_take_bit:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2720
  \<open>of_nat (take_bit n m) = take_bit n (of_nat m)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2721
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2722
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2723
end
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2724
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2725
lemma push_bit_nat_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2726
  \<open>push_bit n (nat k) = nat (push_bit n k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2727
  by (cases \<open>k \<ge> 0\<close>) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2728
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2729
lemma drop_bit_nat_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2730
  \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2731
  apply (cases \<open>k \<ge> 0\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2732
   apply (simp_all add: drop_bit_eq_div nat_div_distrib nat_power_eq not_le)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2733
  apply (simp add: divide_int_def)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2734
  done
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2735
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2736
lemma take_bit_nat_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2737
  \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2738
  using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2739
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2740
lemma nat_take_bit_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2741
  \<open>nat (take_bit n k) = take_bit n (nat k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2742
  if \<open>k \<ge> 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2743
  using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2744
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2745
lemma not_exp_less_eq_0_int [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2746
  \<open>\<not> 2 ^ n \<le> (0::int)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2747
  by (simp add: power_le_zero_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2748
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2749
lemma half_nonnegative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2750
  \<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2751
proof (cases \<open>k \<ge> 0\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2752
  case True
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2753
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2754
    by (auto simp add: divide_int_def sgn_1_pos)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2755
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2756
  case False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2757
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2758
    by (auto simp add: divide_int_def not_le elim!: evenE)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2759
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2760
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2761
lemma half_negative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2762
  \<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2763
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2764
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2765
lemma push_bit_of_Suc_0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2766
  "push_bit n (Suc 0) = 2 ^ n"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2767
  using push_bit_of_1 [where ?'a = nat] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2768
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2769
lemma take_bit_of_Suc_0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2770
  "take_bit n (Suc 0) = of_bool (0 < n)"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2771
  using take_bit_of_1 [where ?'a = nat] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2772
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2773
lemma drop_bit_of_Suc_0 [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2774
  "drop_bit n (Suc 0) = of_bool (n = 0)"
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2775
  using drop_bit_of_1 [where ?'a = nat] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2776
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2777
lemma int_bit_bound:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2778
  fixes k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2779
  obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2780
    and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2781
proof -
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2782
  obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2783
  proof (cases \<open>k \<ge> 0\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2784
    case True
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2785
    moreover from power_gt_expt [of 2 \<open>nat k\<close>]
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2786
    have \<open>nat k < 2 ^ nat k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2787
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2788
    then have \<open>int (nat k) < int (2 ^ nat k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2789
      by (simp only: of_nat_less_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2790
    ultimately have *: \<open>k div 2 ^ nat k = 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2791
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2792
    show thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2793
    proof (rule that [of \<open>nat k\<close>])
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2794
      fix m
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2795
      assume \<open>nat k \<le> m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2796
      then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2797
        by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2798
    qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2799
  next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2800
    case False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2801
    moreover from power_gt_expt [of 2 \<open>nat (- k)\<close>]
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2802
    have \<open>nat (- k) < 2 ^ nat (- k)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2803
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2804
    then have \<open>int (nat (- k)) < int (2 ^ nat (- k))\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2805
      by (simp only: of_nat_less_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2806
    ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2807
      by (subst div_pos_neg_trivial) simp_all
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2808
    then have *: \<open>k div 2 ^ nat (- k) = - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2809
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2810
    show thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2811
    proof (rule that [of \<open>nat (- k)\<close>])
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2812
      fix m
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2813
      assume \<open>nat (- k) \<le> m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2814
      then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2815
        by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2816
    qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2817
  qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2818
  show thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2819
  proof (cases \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2820
    case True
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2821
    then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2822
      by blast
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2823
    with True that [of 0] show thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2824
      by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2825
  next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2826
    case False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2827
    then obtain r where **: \<open>bit k r \<noteq> bit k q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2828
      by blast
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2829
    have \<open>r < q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2830
      by (rule ccontr) (use * [of r] ** in simp)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2831
    define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2832
    moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2833
      using ** N_def \<open>r < q\<close> by auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2834
    moreover define n where \<open>n = Suc (Max N)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2835
    ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2836
      apply auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2837
         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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2838
        apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2839
        apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2840
      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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2841
      done
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2842
    have \<open>bit k (Max N) \<noteq> bit k n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2843
      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)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2844
    show thesis apply (rule that [of n])
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2845
      using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2846
      using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2847
  qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2848
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2849
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2850
context semiring_bit_operations
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2851
begin
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2852
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2853
lemma of_nat_and_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2854
  \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2855
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2856
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2857
lemma of_nat_or_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2858
  \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2859
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2860
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2861
lemma of_nat_xor_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2862
  \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2863
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2864
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2865
end
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2866
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2867
context ring_bit_operations
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2868
begin
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2869
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2870
lemma of_nat_mask_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2871
  \<open>of_nat (mask n) = mask n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2872
  by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2873
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2874
end
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2875
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2876
lemma Suc_mask_eq_exp:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2877
  \<open>Suc (mask n) = 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2878
  by (simp add: mask_eq_exp_minus_1)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2879
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2880
lemma less_eq_mask:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2881
  \<open>n \<le> mask n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2882
  by (simp add: mask_eq_exp_minus_1 le_diff_conv2)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2883
    (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2884
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2885
lemma less_mask:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2886
  \<open>n < mask n\<close> if \<open>Suc 0 < n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2887
proof -
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2888
  define m where \<open>m = n - 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2889
  with that have *: \<open>n = m + 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2890
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2891
  have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2892
    by (induction m) simp_all
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2893
  then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2894
    by (simp add: Suc_mask_eq_exp)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2895
  then have \<open>m + 2 < mask (m + 2)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2896
    by (simp add: less_le)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2897
  with * show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2898
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2899
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2900
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  2901
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2902
subsection \<open>Bit concatenation\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2903
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2904
definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close>
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2905
  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
  2906
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  2907
lemma bit_concat_bit_iff [bit_simps]:
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2908
  \<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
  2909
  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
  2910
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2911
lemma concat_bit_eq:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2912
  \<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
  2913
  by (simp add: concat_bit_def take_bit_eq_mask
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2914
    bit_and_iff bit_mask_iff bit_push_bit_iff disjunctive_add)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2915
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2916
lemma concat_bit_0 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2917
  \<open>concat_bit 0 k l = l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2918
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2919
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2920
lemma concat_bit_Suc:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2921
  \<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
  2922
  by (simp add: concat_bit_eq take_bit_Suc push_bit_double)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2923
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2924
lemma concat_bit_of_zero_1 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2925
  \<open>concat_bit n 0 l = push_bit n l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2926
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2927
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2928
lemma concat_bit_of_zero_2 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2929
  \<open>concat_bit n k 0 = take_bit n k\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2930
  by (simp add: concat_bit_def take_bit_eq_mask)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2931
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2932
lemma concat_bit_nonnegative_iff [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2933
  \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2934
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2935
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2936
lemma concat_bit_negative_iff [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2937
  \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2938
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2939
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2940
lemma concat_bit_assoc:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2941
  \<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
  2942
  by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2943
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2944
lemma concat_bit_assoc_sym:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2945
  \<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
  2946
  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
  2947
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2948
lemma concat_bit_eq_iff:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2949
  \<open>concat_bit n k l = concat_bit n r s
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2950
    \<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
  2951
proof
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2952
  assume ?Q
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2953
  then show ?P
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2954
    by (simp add: concat_bit_def)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2955
next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2956
  assume ?P
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2957
  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
  2958
    by (simp add: bit_eq_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2959
  have \<open>take_bit n k = take_bit n r\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2960
  proof (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2961
    fix m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2962
    from * [of m]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2963
    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
  2964
      by (auto simp add: bit_take_bit_iff bit_concat_bit_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2965
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2966
  moreover have \<open>push_bit n l = push_bit n s\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2967
  proof (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2968
    fix m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2969
    from * [of m]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2970
    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
  2971
      by (auto simp add: bit_push_bit_iff bit_concat_bit_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2972
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2973
  then have \<open>l = s\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2974
    by (simp add: push_bit_eq_mult)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2975
  ultimately show ?Q
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2976
    by (simp add: concat_bit_def)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2977
qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2978
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2979
lemma take_bit_concat_bit_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2980
  \<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
  2981
  by (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2982
    (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def)  
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  2983
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2984
lemma concat_bit_take_bit_eq:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2985
  \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2986
  by (simp add: concat_bit_def [abs_def])
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2987
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2988
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2989
subsection \<open>Taking bits with sign propagation\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  2990
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2991
context ring_bit_operations
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2992
begin
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  2993
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2994
definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2995
  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
  2996
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2997
lemma signed_take_bit_eq_if_positive:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  2998
  \<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
  2999
  using that by (simp add: signed_take_bit_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3000
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3001
lemma signed_take_bit_eq_if_negative:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3002
  \<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
  3003
  using that by (simp add: signed_take_bit_def)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3004
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3005
lemma even_signed_take_bit_iff:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3006
  \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3007
  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
  3008
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  3009
lemma bit_signed_take_bit_iff [bit_simps]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3010
  \<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
  3011
  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
  3012
    (use exp_eq_0_imp_not_bit in blast)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3013
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3014
lemma signed_take_bit_0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3015
  \<open>signed_take_bit 0 a = - (a mod 2)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3016
  by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3017
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3018
lemma signed_take_bit_Suc:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3019
  \<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
  3020
proof (rule bit_eqI)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3021
  fix m
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3022
  assume *: \<open>2 ^ m \<noteq> 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3023
  show \<open>bit (signed_take_bit (Suc n) a) m \<longleftrightarrow>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3024
    bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3025
  proof (cases m)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3026
    case 0
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3027
    then show ?thesis
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3028
      by (simp add: even_signed_take_bit_iff)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3029
  next
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3030
    case (Suc m)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3031
    with * have \<open>2 ^ m \<noteq> 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3032
      by (metis mult_not_zero power_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3033
    with Suc show ?thesis
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3034
      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
  3035
        ac_simps flip: bit_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3036
  qed
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3037
qed
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3038
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3039
lemma signed_take_bit_of_0 [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3040
  \<open>signed_take_bit n 0 = 0\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3041
  by (simp add: signed_take_bit_def)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3042
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3043
lemma signed_take_bit_of_minus_1 [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3044
  \<open>signed_take_bit n (- 1) = - 1\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3045
  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
  3046
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3047
lemma signed_take_bit_Suc_1 [simp]:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3048
  \<open>signed_take_bit (Suc n) 1 = 1\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3049
  by (simp add: signed_take_bit_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3050
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3051
lemma signed_take_bit_rec:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3052
  \<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
  3053
  by (cases n) (simp_all add: signed_take_bit_Suc)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3054
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3055
lemma signed_take_bit_eq_iff_take_bit_eq:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3056
  \<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
  3057
proof -
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3058
  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
  3059
    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
  3060
      (use exp_eq_0_imp_not_bit in fastforce)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3061
  then show ?thesis
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3062
    by (simp add: bit_eq_iff fun_eq_iff)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3063
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3064
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3065
lemma signed_take_bit_signed_take_bit [simp]:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3066
  \<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
  3067
proof (rule bit_eqI)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3068
  fix q
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3069
  show \<open>bit (signed_take_bit m (signed_take_bit n a)) q \<longleftrightarrow>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3070
    bit (signed_take_bit (min m n) a) q\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3071
    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
  3072
      (use le_Suc_ex exp_add_not_zero_imp in blast)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3073
qed
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3074
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3075
lemma signed_take_bit_take_bit:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3076
  \<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
  3077
  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
  3078
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3079
lemma take_bit_signed_take_bit:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3080
  \<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
  3081
  using that by (rule le_SucE; intro bit_eqI)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3082
   (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
  3083
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3084
end
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3085
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3086
text \<open>Modulus centered around 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3087
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3088
lemma signed_take_bit_eq_concat_bit:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3089
  \<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
  3090
  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
  3091
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3092
lemma signed_take_bit_add:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3093
  \<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
  3094
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3095
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3096
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3097
     (take_bit (Suc n) (signed_take_bit n k) +
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3098
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3099
    take_bit (Suc n) (k + l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3100
    by (simp add: take_bit_signed_take_bit take_bit_add)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3101
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3102
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3103
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3104
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3105
lemma signed_take_bit_diff:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3106
  \<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
  3107
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3108
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3109
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3110
     (take_bit (Suc n) (signed_take_bit n k) -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3111
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3112
    take_bit (Suc n) (k - l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3113
    by (simp add: take_bit_signed_take_bit take_bit_diff)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3114
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3115
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3116
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3117
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3118
lemma signed_take_bit_minus:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3119
  \<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
  3120
  for k :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3121
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3122
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3123
     (- take_bit (Suc n) (signed_take_bit n k)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3124
    take_bit (Suc n) (- k)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3125
    by (simp add: take_bit_signed_take_bit take_bit_minus)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3126
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3127
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3128
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3129
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3130
lemma signed_take_bit_mult:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3131
  \<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
  3132
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3133
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3134
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3135
     (take_bit (Suc n) (signed_take_bit n k) *
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3136
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3137
    take_bit (Suc n) (k * l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3138
    by (simp add: take_bit_signed_take_bit take_bit_mult)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3139
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3140
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3141
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3142
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3143
lemma signed_take_bit_eq_take_bit_minus:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3144
  \<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
  3145
  for k :: int
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3146
proof (cases \<open>bit k n\<close>)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3147
  case True
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3148
  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
  3149
    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
  3150
  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
  3151
    by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3152
  with True show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3153
    by (simp flip: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3154
next
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3155
  case False
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3156
  show ?thesis
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3157
    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
  3158
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3159
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3160
lemma signed_take_bit_eq_take_bit_shift:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3161
  \<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
  3162
  for k :: int
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3163
proof -
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3164
  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
  3165
    by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3166
  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
  3167
    by (simp add: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3168
  also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3169
    by (rule disjunctive_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3170
      (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3171
  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
  3172
  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
  3173
    by (simp only: take_bit_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3174
  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
  3175
    by (simp add: take_bit_Suc_from_most)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3176
  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
  3177
    by (simp add: ac_simps)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3178
  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
  3179
    by (rule disjunctive_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3180
      (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3181
  finally show ?thesis
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3182
    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
  3183
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3184
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3185
lemma signed_take_bit_nonnegative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3186
  \<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
  3187
  for k :: int
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3188
  by (simp add: signed_take_bit_def not_less concat_bit_def)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3189
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3190
lemma signed_take_bit_negative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3191
  \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3192
  for k :: int
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3193
  by (simp add: signed_take_bit_def not_less concat_bit_def)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3194
73868
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3195
lemma signed_take_bit_int_greater_eq_minus_exp [simp]:
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3196
  \<open>- (2 ^ n) \<le> signed_take_bit n k\<close>
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3197
  for k :: int
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3198
  by (simp add: signed_take_bit_eq_take_bit_shift)
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3199
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3200
lemma signed_take_bit_int_less_exp [simp]:
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3201
  \<open>signed_take_bit n k < 2 ^ n\<close>
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3202
  for k :: int
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3203
  using take_bit_int_less_exp [of \<open>Suc n\<close>]
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3204
  by (simp add: signed_take_bit_eq_take_bit_shift)
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3205
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3206
lemma signed_take_bit_int_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3207
  \<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
  3208
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3209
  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
  3210
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3211
lemma signed_take_bit_int_eq_self:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3212
  \<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
  3213
  for k :: int
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3214
  using that by (simp add: signed_take_bit_int_eq_self_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3215
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3216
lemma signed_take_bit_int_less_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3217
  \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3218
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3219
  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
  3220
    linarith
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3221
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3222
lemma signed_take_bit_int_less_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3223
  \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3224
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3225
  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
  3226
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3227
lemma signed_take_bit_int_greater_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3228
  \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3229
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3230
  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
  3231
    linarith
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3232
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3233
lemma signed_take_bit_int_greater_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3234
  \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3235
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3236
  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
  3237
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3238
lemma signed_take_bit_int_greater_eq:
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3239
  \<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
  3240
  for k :: int
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3241
  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
  3242
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3243
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3244
lemma signed_take_bit_int_less_eq:
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3245
  \<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
  3246
  for k :: int
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3247
  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
  3248
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3249
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3250
lemma signed_take_bit_Suc_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3251
  \<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
  3252
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3253
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3254
lemma signed_take_bit_Suc_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3255
  \<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
  3256
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3257
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3258
lemma signed_take_bit_Suc_minus_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3259
  \<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
  3260
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3261
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3262
lemma signed_take_bit_Suc_minus_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3263
  \<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
  3264
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3265
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3266
lemma signed_take_bit_numeral_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3267
  \<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
  3268
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3269
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3270
lemma signed_take_bit_numeral_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3271
  \<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
  3272
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3273
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3274
lemma signed_take_bit_numeral_minus_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3275
  \<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
  3276
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3277
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3278
lemma signed_take_bit_numeral_minus_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3279
  \<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
  3280
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3281
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3282
lemma signed_take_bit_code [code]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3283
  \<open>signed_take_bit n a =
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3284
  (let l = take_bit (Suc n) a
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3285
   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
  3286
proof -
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3287
  have *: \<open>take_bit (Suc n) a + push_bit n (- 2) =
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3288
    take_bit (Suc n) a OR NOT (mask (Suc n))\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3289
    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
  3290
       simp flip: push_bit_minus_one_eq_not_mask)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3291
  show ?thesis
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3292
    by (rule bit_eqI)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3293
      (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
  3294
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3295
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3296
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  3297
subsection \<open>Horner sums\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  3298
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3299
context semiring_bit_operations
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3300
begin
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3301
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3302
lemma horner_sum_bit_eq_take_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3303
  \<open>horner_sum of_bool 2 (map (bit a) [0..<n]) = take_bit n a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3304
proof (induction a arbitrary: n rule: bits_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3305
  case (stable a)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3306
  moreover have \<open>bit a = (\<lambda>_. odd a)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3307
    using stable by (simp add: stable_imp_bit_iff_odd fun_eq_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3308
  moreover have \<open>{q. q < n} = {0..<n}\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3309
    by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3310
  ultimately show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3311
    by (simp add: stable_imp_take_bit_eq horner_sum_eq_sum mask_eq_sum_exp)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3312
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3313
  case (rec a b)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3314
  show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3315
  proof (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3316
    case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3317
    then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3318
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3319
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3320
    case (Suc m)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3321
    have \<open>map (bit (of_bool b + 2 * a)) [0..<Suc m] = b # map (bit (of_bool b + 2 * a)) [Suc 0..<Suc m]\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3322
      by (simp only: upt_conv_Cons) simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3323
    also have \<open>\<dots> = b # map (bit a) [0..<m]\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3324
      by (simp only: flip: map_Suc_upt) (simp add: bit_Suc rec.hyps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3325
    finally show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3326
      using Suc rec.IH [of m] by (simp add: take_bit_Suc rec.hyps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3327
        (simp_all add: ac_simps mod_2_eq_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3328
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3329
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3330
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3331
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3332
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  3333
context unique_euclidean_semiring_with_bit_operations
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3334
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3335
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3336
lemma bit_horner_sum_bit_iff [bit_simps]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3337
  \<open>bit (horner_sum of_bool 2 bs) n \<longleftrightarrow> n < length bs \<and> bs ! n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3338
proof (induction bs arbitrary: n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3339
  case Nil
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3340
  then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3341
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3342
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3343
  case (Cons b bs)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3344
  show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3345
  proof (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3346
    case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3347
    then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3348
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3349
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3350
    case (Suc m)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3351
    with bit_rec [of _ n] Cons.prems Cons.IH [of m]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3352
    show ?thesis by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3353
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3354
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3355
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3356
lemma take_bit_horner_sum_bit_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3357
  \<open>take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3358
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3359
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3360
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3361
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3362
lemma horner_sum_of_bool_2_less:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3363
  \<open>(horner_sum of_bool 2 bs :: int) < 2 ^ length bs\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3364
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3365
  have \<open>(\<Sum>n = 0..<length bs. of_bool (bs ! n) * (2::int) ^ n) \<le> (\<Sum>n = 0..<length bs. 2 ^ n)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3366
    by (rule sum_mono) simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3367
  also have \<open>\<dots> = 2 ^ length bs - 1\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3368
    by (induction bs) simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3369
  finally show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3370
    by (simp add: horner_sum_eq_sum)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3371
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3372
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3373
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3374
subsection \<open>Symbolic computations on numeral expressions\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3375
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3376
fun and_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3377
where
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3378
  \<open>and_num num.One num.One = Some num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3379
| \<open>and_num num.One (num.Bit0 n) = None\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3380
| \<open>and_num num.One (num.Bit1 n) = Some num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3381
| \<open>and_num (num.Bit0 m) num.One = None\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3382
| \<open>and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3383
| \<open>and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3384
| \<open>and_num (num.Bit1 m) num.One = Some num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3385
| \<open>and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3386
| \<open>and_num (num.Bit1 m) (num.Bit1 n) = (case and_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3387
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3388
fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3389
where
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3390
  \<open>and_not_num num.One num.One = None\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3391
| \<open>and_not_num num.One (num.Bit0 n) = Some num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3392
| \<open>and_not_num num.One (num.Bit1 n) = None\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3393
| \<open>and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3394
| \<open>and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3395
| \<open>and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3396
| \<open>and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3397
| \<open>and_not_num (num.Bit1 m) (num.Bit0 n) = (case and_not_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3398
| \<open>and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3399
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3400
fun or_num :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3401
where
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3402
  \<open>or_num num.One num.One = num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3403
| \<open>or_num num.One (num.Bit0 n) = num.Bit1 n\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3404
| \<open>or_num num.One (num.Bit1 n) = num.Bit1 n\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3405
| \<open>or_num (num.Bit0 m) num.One = num.Bit1 m\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3406
| \<open>or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3407
| \<open>or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3408
| \<open>or_num (num.Bit1 m) num.One = num.Bit1 m\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3409
| \<open>or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3410
| \<open>or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3411
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3412
fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3413
where
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3414
  \<open>or_not_num_neg num.One num.One = num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3415
| \<open>or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3416
| \<open>or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3417
| \<open>or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3418
| \<open>or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3419
| \<open>or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3420
| \<open>or_not_num_neg (num.Bit1 n) num.One = num.One\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3421
| \<open>or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3422
| \<open>or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3423
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3424
fun xor_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3425
where
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3426
  \<open>xor_num num.One num.One = None\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3427
| \<open>xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3428
| \<open>xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3429
| \<open>xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3430
| \<open>xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3431
| \<open>xor_num (num.Bit0 m) (num.Bit1 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3432
| \<open>xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3433
| \<open>xor_num (num.Bit1 m) (num.Bit0 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3434
| \<open>xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3435
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3436
lemma int_numeral_and_num:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3437
  \<open>numeral m AND numeral n = (case and_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3438
  by (induction m n rule: and_num.induct) (simp_all split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3439
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3440
lemma and_num_eq_None_iff:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3441
  \<open>and_num m n = None \<longleftrightarrow> numeral m AND numeral n = (0::int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3442
  by (simp add: int_numeral_and_num split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3443
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3444
lemma and_num_eq_Some_iff:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3445
  \<open>and_num m n = Some q \<longleftrightarrow> numeral m AND numeral n = (numeral q :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3446
  by (simp add: int_numeral_and_num split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3447
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3448
lemma int_numeral_and_not_num:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3449
  \<open>numeral m AND NOT (numeral n) = (case and_not_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3450
  by (induction m n rule: and_not_num.induct) (simp_all add: add_One BitM_inc_eq not_int_def split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3451
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3452
lemma int_numeral_not_and_num:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3453
  \<open>NOT (numeral m) AND numeral n = (case and_not_num n m of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3454
  using int_numeral_and_not_num [of n m] by (simp add: ac_simps)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3455
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3456
lemma and_not_num_eq_None_iff:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3457
  \<open>and_not_num m n = None \<longleftrightarrow> numeral m AND NOT (numeral n) = (0::int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3458
  by (simp add: int_numeral_and_not_num split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3459
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3460
lemma and_not_num_eq_Some_iff:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3461
  \<open>and_not_num m n = Some q \<longleftrightarrow> numeral m AND NOT (numeral n) = (numeral q :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3462
  by (simp add: int_numeral_and_not_num split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3463
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3464
lemma int_numeral_or_num:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3465
  \<open>numeral m OR numeral n = (numeral (or_num m n) :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3466
  by (induction m n rule: or_num.induct) simp_all
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3467
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3468
lemma numeral_or_num_eq:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3469
  \<open>numeral (or_num m n) = (numeral m OR numeral n :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3470
  by (simp add: int_numeral_or_num)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3471
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3472
lemma int_numeral_or_not_num_neg:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3473
  \<open>numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3474
  by (induction m n rule: or_not_num_neg.induct) (simp_all add: add_One BitM_inc_eq not_int_def)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3475
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3476
lemma int_numeral_not_or_num_neg:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3477
  \<open>NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3478
  using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3479
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3480
lemma numeral_or_not_num_eq:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3481
  \<open>numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3482
  using int_numeral_or_not_num_neg [of m n] by simp
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3483
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3484
lemma int_numeral_xor_num:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3485
  \<open>numeral m XOR numeral n = (case xor_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3486
  by (induction m n rule: xor_num.induct) (simp_all split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3487
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3488
lemma xor_num_eq_None_iff:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3489
  \<open>xor_num m n = None \<longleftrightarrow> numeral m XOR numeral n = (0::int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3490
  by (simp add: int_numeral_xor_num split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3491
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3492
lemma xor_num_eq_Some_iff:
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3493
  \<open>xor_num m n = Some q \<longleftrightarrow> numeral m XOR numeral n = (numeral q :: int)\<close>
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3494
  by (simp add: int_numeral_xor_num split: option.split)
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3495
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  3496
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3497
subsection \<open>Key ideas of bit operations\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3498
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3499
text \<open>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3500
  When formalizing bit operations, it is tempting to represent
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3501
  bit values as explicit lists over a binary type. This however
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3502
  is a bad idea, mainly due to the inherent ambiguities in
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3503
  representation concerning repeating leading bits.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3504
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3505
  Hence this approach avoids such explicit lists altogether
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3506
  following an algebraic path:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3507
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3508
  \<^item> Bit values are represented by numeric types: idealized
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3509
    unbounded bit values can be represented by type \<^typ>\<open>int\<close>,
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3510
    bounded bit values by quotient types over \<^typ>\<open>int\<close>.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3511
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3512
  \<^item> (A special case are idealized unbounded bit values ending
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3513
    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
  3514
    only support a restricted set of operations).
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3515
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3516
  \<^item> From this idea follows that
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3517
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3518
      \<^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
  3519
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3520
      \<^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
  3521
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3522
  \<^item> Concerning bounded bit values, iterated shifts to the left
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3523
    may result in eliminating all bits by shifting them all
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3524
    beyond the boundary.  The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3525
    represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3526
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
  3527
  \<^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
  3528
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3529
  \<^item> This leads to the most fundamental properties of bit values:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3530
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3531
      \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3532
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3533
      \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3534
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3535
  \<^item> Typical operations are characterized as follows:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3536
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3537
      \<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3538
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  3539
      \<^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
  3540
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3541
      \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3542
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3543
      \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3544
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3545
      \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3546
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3547
      \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3548
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3549
      \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3550
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3551
      \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3552
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3553
      \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3554
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3555
      \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3556
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3557
      \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3558
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3559
      \<^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
  3560
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3561
      \<^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
  3562
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3563
      \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]}
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3564
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3565
      \<^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
  3566
\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3567
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  3568
find_theorems \<open>(_ AND _) * _ = _\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  3569
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3570
no_notation
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3571
  "and"  (infixr \<open>AND\<close> 64)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3572
    and or  (infixr \<open>OR\<close>  59)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3573
    and xor  (infixr \<open>XOR\<close> 59)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3574
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3575
bundle bit_operations_syntax
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  3576
begin
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3577
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3578
notation
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3579
  "and"  (infixr \<open>AND\<close> 64)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3580
    and or  (infixr \<open>OR\<close>  59)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3581
    and xor  (infixr \<open>XOR\<close> 59)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3582
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  3583
end
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3584
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  3585
end