src/HOL/Bit_Operations.thy
author haftmann
Sat, 19 Apr 2025 17:39:06 +0200
changeset 82524 df5b2785abd6
parent 82289 26fbbf43863b
permissions -rw-r--r--
more lemmas
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
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
     8
begin
74101
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
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 79531
diff changeset
    12
class semiring_bits = semiring_parity + semiring_modulo_trivial +
79480
c7cb1bf6efa0 consolidated name of lemma analogously to nat/int/word_bit_induct
haftmann
parents: 79117
diff changeset
    13
  assumes bit_induct [case_names stable rec]:
74101
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>
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
    17
  assumes bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
    18
    and half_div_exp_eq: \<open>a div 2 div 2 ^ n = a div 2 ^ Suc n\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    19
    and even_double_div_exp_iff: \<open>2 ^ Suc n \<noteq> 0 \<Longrightarrow> even (2 * a div 2 ^ Suc n) \<longleftrightarrow> even (a div 2 ^ n)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    20
  fixes bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    21
  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
    22
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    23
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    24
text \<open>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    25
  Having \<^const>\<open>bit\<close> as definitional class operation
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    26
  takes into account that specific instances can be implemented
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    27
  differently wrt. code generation.
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    28
\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    29
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    30
lemma half_1 [simp]:
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    31
  \<open>1 div 2 = 0\<close>
79488
62d8c6c08fb2 consolidated lemma name
haftmann
parents: 79481
diff changeset
    32
  using even_half_succ_eq [of 0] by simp
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    33
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    34
lemma div_exp_eq_funpow_half:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    35
  \<open>a div 2 ^ n = ((\<lambda>a. a div 2) ^^ n) a\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    36
proof -
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    37
  have \<open>((\<lambda>a. a div 2) ^^ n) = (\<lambda>a. a div 2 ^ n)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    38
    by (induction n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    39
      (simp_all del: funpow.simps power.simps add: power_0 funpow_Suc_right half_div_exp_eq)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    40
  then show ?thesis
79481
8205977e9e2c simplified specification of type class
haftmann
parents: 79480
diff changeset
    41
    by simp
8205977e9e2c simplified specification of type class
haftmann
parents: 79480
diff changeset
    42
qed
8205977e9e2c simplified specification of type class
haftmann
parents: 79480
diff changeset
    43
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    44
lemma div_exp_eq:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    45
  \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    46
  by (simp add: div_exp_eq_funpow_half Groups.add.commute [of m] funpow_add)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
    47
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
    48
lemma bit_0:
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    49
  \<open>bit a 0 \<longleftrightarrow> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    50
  by (simp add: bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    51
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    52
lemma bit_Suc:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    53
  \<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
    54
  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
    55
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    56
lemma bit_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    57
  \<open>bit a n \<longleftrightarrow> (if n = 0 then odd a else bit (a div 2) (n - 1))\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
    58
  by (cases n) (simp_all add: bit_Suc bit_0)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    59
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    60
context
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    61
  fixes a
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    62
  assumes stable: \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    63
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    64
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    65
lemma bits_stable_imp_add_self:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    66
  \<open>a + a mod 2 = 0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    67
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    68
  have \<open>a div 2 * 2 + a mod 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    69
    by (fact div_mult_mod_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    70
  then have \<open>a * 2 + a mod 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    71
    by (simp add: stable)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    72
  then show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    73
    by (simp add: mult_2_right ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    74
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    75
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    76
lemma stable_imp_bit_iff_odd:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    77
  \<open>bit a n \<longleftrightarrow> odd a\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
    78
  by (induction n) (simp_all add: stable bit_Suc bit_0)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    79
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    80
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    81
79585
dafb3d343cd6 more lemmas and more correct lemma names
haftmann
parents: 79555
diff changeset
    82
lemma bit_iff_odd_imp_stable:
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    83
  \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close>
79480
c7cb1bf6efa0 consolidated name of lemma analogously to nat/int/word_bit_induct
haftmann
parents: 79117
diff changeset
    84
using that proof (induction a rule: bit_induct)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    85
  case (stable a)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    86
  then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    87
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    88
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    89
  case (rec a b)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    90
  from rec.prems [of 1] have [simp]: \<open>b = odd a\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
    91
    by (simp add: rec.hyps bit_Suc bit_0)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    92
  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
    93
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    94
  have \<open>bit a n \<longleftrightarrow> odd a\<close> for n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    95
    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
    96
  then have \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    97
    by (rule rec.IH)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
    98
  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
    99
    by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   100
  also have \<open>\<dots> = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   101
    using mult_div_mod_eq [of 2 a]
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   102
    by (simp add: of_bool_odd_eq_mod_2)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   103
  finally show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   104
    using \<open>a div 2 = a\<close> by (simp add: hyp)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   105
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   106
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   107
lemma even_succ_div_exp [simp]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   108
  \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   109
proof (cases n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   110
  case 0
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   111
  with that show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   112
    by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   113
next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   114
  case (Suc n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   115
  with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   116
  proof (induction n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   117
    case 0
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   118
    then show ?case
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   119
      by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   120
  next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   121
    case (Suc n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   122
    then show ?case
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   123
      using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric]
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   124
      by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   125
  qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   126
  with Suc show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   127
    by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   128
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   129
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   130
lemma even_succ_mod_exp [simp]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   131
  \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   132
  using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] div_mult_mod_eq [of a \<open>2 ^ n\<close>] that
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   133
  by simp (metis (full_types) add.left_commute add_left_imp_eq)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   134
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   135
lemma half_numeral_Bit1_eq [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   136
  \<open>numeral (num.Bit1 m) div 2 = numeral (num.Bit0 m) div 2\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   137
  using even_half_succ_eq [of \<open>2 * numeral m\<close>]
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   138
  by simp
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   139
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   140
lemma double_half_numeral_Bit_0_eq [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   141
  \<open>2 * (numeral (num.Bit0 m) div 2) = numeral (num.Bit0 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   142
  \<open>(numeral (num.Bit0 m) div 2) * 2 = numeral (num.Bit0 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   143
  using mod_mult_div_eq [of \<open>numeral (Num.Bit0 m)\<close> 2]
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   144
  by (simp_all add: mod2_eq_if ac_simps)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
   145
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   146
named_theorems bit_simps \<open>Simplification rules for \<^const>\<open>bit\<close>\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   147
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   148
definition possible_bit :: \<open>'a itself \<Rightarrow> nat \<Rightarrow> bool\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   149
  where \<open>possible_bit TYPE('a) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
79018
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
   150
  \<comment> \<open>This auxiliary avoids non-termination with extensionality.\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   151
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   152
lemma possible_bit_0 [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   153
  \<open>possible_bit TYPE('a) 0\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   154
  by (simp add: possible_bit_def)
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   155
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   156
lemma fold_possible_bit:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   157
  \<open>2 ^ n = 0 \<longleftrightarrow> \<not> possible_bit TYPE('a) n\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   158
  by (simp add: possible_bit_def)
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   159
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   160
lemma bit_imp_possible_bit:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   161
  \<open>possible_bit TYPE('a) n\<close> if \<open>bit a n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   162
  by (rule ccontr) (use that in \<open>auto simp: bit_iff_odd possible_bit_def\<close>)
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   163
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   164
lemma impossible_bit:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   165
  \<open>\<not> bit a n\<close> if \<open>\<not> possible_bit TYPE('a) n\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   166
  using that by (blast dest: bit_imp_possible_bit)
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   167
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   168
lemma possible_bit_less_imp:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   169
  \<open>possible_bit TYPE('a) j\<close> if \<open>possible_bit TYPE('a) i\<close> \<open>j \<le> i\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   170
  using power_add [of 2 j \<open>i - j\<close>] that mult_not_zero [of \<open>2 ^ j\<close> \<open>2 ^ (i - j)\<close>]
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   171
  by (simp add: possible_bit_def)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   172
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   173
lemma possible_bit_min [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   174
  \<open>possible_bit TYPE('a) (min i j) \<longleftrightarrow> possible_bit TYPE('a) i \<or> possible_bit TYPE('a) j\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   175
  by (auto simp: min_def elim: possible_bit_less_imp)
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   176
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   177
lemma bit_eqI:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   178
  \<open>a = b\<close> if \<open>\<And>n. possible_bit TYPE('a) n \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   179
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   180
  have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   181
  proof (cases \<open>possible_bit TYPE('a) n\<close>)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   182
    case False
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   183
    then show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   184
      by (simp add: impossible_bit)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   185
  next
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   186
    case True
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   187
    then show ?thesis
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   188
      by (rule that)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   189
  qed
79480
c7cb1bf6efa0 consolidated name of lemma analogously to nat/int/word_bit_induct
haftmann
parents: 79117
diff changeset
   190
  then show ?thesis proof (induction a arbitrary: b rule: bit_induct)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   191
    case (stable a)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   192
    from stable(2) [of 0] have **: \<open>even b \<longleftrightarrow> even a\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
   193
      by (simp add: bit_0)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   194
    have \<open>b div 2 = b\<close>
79585
dafb3d343cd6 more lemmas and more correct lemma names
haftmann
parents: 79555
diff changeset
   195
    proof (rule bit_iff_odd_imp_stable)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   196
      fix n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   197
      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
   198
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   199
      also have \<open>bit a n \<longleftrightarrow> odd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   200
        using stable by (simp add: stable_imp_bit_iff_odd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   201
      finally show \<open>bit b n \<longleftrightarrow> odd b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   202
        by (simp add: **)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   203
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   204
    from ** have \<open>a mod 2 = b mod 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   205
      by (simp add: mod2_eq_if)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   206
    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
   207
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   208
    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
   209
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   210
    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
   211
      by (simp add: bits_stable_imp_add_self)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   212
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   213
    case (rec a p)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   214
    from rec.prems [of 0] have [simp]: \<open>p = odd b\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
   215
      by (simp add: bit_0)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   216
    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
   217
      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
   218
    then have \<open>a = b div 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   219
      by (rule rec.IH)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   220
    then have \<open>2 * a = 2 * (b div 2)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   221
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   222
    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
   223
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   224
    also have \<open>\<dots> = b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   225
      by (fact mod_mult_div_eq)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   226
    finally show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   227
      by (auto simp: mod2_eq_if)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   228
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   229
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   230
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   231
lemma bit_eq_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   232
  \<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
   233
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   234
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   235
  then show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   236
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   237
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   238
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   239
  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
   240
    by simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   241
  show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   242
  proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   243
    fix n
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   244
    show \<open>bit a n \<longleftrightarrow> bit b n\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   245
    proof (cases n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   246
      case 0
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   247
      with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
   248
        by (simp add: bit_0)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   249
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   250
      case (Suc n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   251
      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
   252
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   253
      ultimately show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   254
        by (simp add: bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   255
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   256
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   257
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   258
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   259
lemma bit_eq_iff:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   260
  \<open>a = b \<longleftrightarrow> (\<forall>n. possible_bit TYPE('a) n \<longrightarrow> bit a n \<longleftrightarrow> bit b n)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   261
  by (auto intro: bit_eqI simp add: possible_bit_def)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   262
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   263
lemma bit_0_eq [simp]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   264
  \<open>bit 0 = \<bottom>\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   265
proof -
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   266
  have \<open>0 div 2 ^ n = 0\<close> for n
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   267
    unfolding div_exp_eq_funpow_half by (induction n) simp_all
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   268
  then show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   269
    by (simp add: fun_eq_iff bit_iff_odd)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   270
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   271
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   272
lemma bit_double_Suc_iff:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   273
  \<open>bit (2 * a) (Suc n) \<longleftrightarrow> possible_bit TYPE('a) (Suc n) \<and> bit a n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   274
  using even_double_div_exp_iff [of n a]
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   275
  by (cases \<open>possible_bit TYPE('a) (Suc n)\<close>)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   276
    (auto simp: bit_iff_odd possible_bit_def)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   277
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   278
lemma bit_double_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   279
  \<open>bit (2 * a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<noteq> 0 \<and> bit a (n - 1)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   280
  by (cases n) (simp_all add: bit_0 bit_double_Suc_iff)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   281
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   282
lemma even_bit_succ_iff:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   283
  \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   284
  using that by (cases \<open>n = 0\<close>) (simp_all add: bit_iff_odd)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   285
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   286
lemma odd_bit_iff_bit_pred:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   287
  \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   288
proof -
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   289
  from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> ..
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   290
  moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   291
    using even_bit_succ_iff by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   292
  ultimately show ?thesis by (simp add: ac_simps)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   293
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   294
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   295
lemma bit_exp_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   296
  \<open>bit (2 ^ m) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n = m\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   297
proof (cases \<open>possible_bit TYPE('a) n\<close>)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   298
  case False
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   299
  then show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   300
    by (simp add: impossible_bit)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   301
next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   302
  case True
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   303
  then show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   304
  proof (induction n arbitrary: m)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   305
    case 0
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   306
    show ?case
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   307
      by (simp add: bit_0)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   308
  next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   309
    case (Suc n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   310
    then have \<open>possible_bit TYPE('a) n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   311
      by (simp add: possible_bit_less_imp)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   312
    show ?case
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   313
    proof (cases m)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   314
      case 0
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   315
      then show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   316
        by (simp add: bit_Suc)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   317
    next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   318
      case (Suc m)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   319
      with Suc.IH [of m] \<open>possible_bit TYPE('a) n\<close> show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   320
        by (simp add: bit_double_Suc_iff)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   321
    qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   322
  qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   323
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   324
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   325
lemma bit_1_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   326
  \<open>bit 1 n \<longleftrightarrow> n = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   327
  using bit_exp_iff [of 0 n] by auto
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   328
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   329
lemma bit_2_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   330
  \<open>bit 2 n \<longleftrightarrow> possible_bit TYPE('a) 1 \<and> n = 1\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   331
  using bit_exp_iff [of 1 n] by auto
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   332
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   333
lemma bit_of_bool_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   334
  \<open>bit (of_bool b) n \<longleftrightarrow> n = 0 \<and> b\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   335
  by (simp add: bit_1_iff)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   336
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   337
lemma bit_mod_2_iff [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   338
  \<open>bit (a mod 2) n \<longleftrightarrow> n = 0 \<and> odd a\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   339
  by (simp add: mod_2_eq_odd bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   340
82289
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   341
lemma stable_index:
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   342
  obtains m where \<open>possible_bit TYPE('a) m\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   343
    \<open>\<And>n. possible_bit TYPE('a) n \<Longrightarrow> n \<ge> m \<Longrightarrow> bit a n \<longleftrightarrow> bit a m\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   344
proof -
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   345
  have \<open>\<exists>m. possible_bit TYPE('a) m \<and> (\<forall>n\<ge>m. possible_bit TYPE('a) n \<longrightarrow> bit a n \<longleftrightarrow> bit a m)\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   346
  proof (induction a rule: bit_induct)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   347
    case (stable a)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   348
    show ?case
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   349
      by (rule exI [of _ \<open>0::nat\<close>]) (simp add: stable_imp_bit_iff_odd stable)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   350
  next
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   351
    case (rec a b)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   352
    then obtain m where \<open>possible_bit TYPE('a) m\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   353
       and hyp: \<open>\<And>n. possible_bit TYPE('a) n \<Longrightarrow> n \<ge> m \<Longrightarrow> bit a n \<longleftrightarrow> bit a m\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   354
      by blast
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   355
    show ?case
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   356
    proof (cases \<open>possible_bit TYPE('a) (Suc m)\<close>)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   357
      case True
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   358
      moreover have \<open>bit (of_bool b + 2 * a) n \<longleftrightarrow> bit (of_bool b + 2 * a) (Suc m)\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   359
        if \<open>possible_bit TYPE('a) n\<close> \<open>Suc m \<le> n\<close> for n
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   360
        using hyp [of \<open>n - 1\<close>] possible_bit_less_imp [of n \<open>n - 1\<close>] rec.hyps that
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   361
        by (cases n) (simp_all add: bit_Suc)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   362
      ultimately show ?thesis
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   363
        by blast
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   364
    next
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   365
      case False
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   366
      have \<open>bit (of_bool b + 2 * a) n \<longleftrightarrow> bit (of_bool b + 2 * a) m\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   367
        if \<open>possible_bit TYPE('a) n\<close> \<open>m \<le> n\<close> for n
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   368
      proof (cases \<open>m = n\<close>)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   369
        case True
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   370
        then show ?thesis
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   371
          by simp
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   372
      next
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   373
        case False
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   374
        with \<open>m \<le> n\<close> have \<open>m < n\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   375
          by simp
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   376
        with \<open>\<not> possible_bit TYPE('a) (Suc m)\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   377
        have \<open>\<not> possible_bit TYPE('a) n\<close> using possible_bit_less_imp [of n \<open>Suc m\<close>]
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   378
          by auto
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   379
        with \<open>possible_bit TYPE('a) n\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   380
        show ?thesis
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   381
          by simp
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   382
      qed
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   383
      with \<open>possible_bit TYPE('a) m\<close> show ?thesis
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   384
        by blast
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   385
    qed
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   386
  qed
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   387
  with that show thesis
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   388
    by blast
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   389
qed
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   390
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
   391
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   392
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   393
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   394
lemma nat_bit_induct [case_names zero even odd]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   395
  \<open>P n\<close> if zero: \<open>P 0\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   396
    and even: \<open>\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   397
    and odd: \<open>\<And>n. P n \<Longrightarrow> P (Suc (2 * n))\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   398
proof (induction n rule: less_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   399
  case (less n)
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   400
  show \<open>P n\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   401
  proof (cases \<open>n = 0\<close>)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   402
    case True with zero show ?thesis by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   403
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   404
    case False
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   405
    with less have hyp: \<open>P (n div 2)\<close> by simp
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   406
    show ?thesis
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   407
    proof (cases \<open>even n\<close>)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   408
      case True
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   409
      then have \<open>n \<noteq> 1\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   410
        by auto
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   411
      with \<open>n \<noteq> 0\<close> have \<open>n div 2 > 0\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   412
        by simp
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   413
      with \<open>even n\<close> hyp even [of \<open>n div 2\<close>] show ?thesis
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   414
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   415
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   416
      case False
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   417
      with hyp odd [of \<open>n div 2\<close>] show ?thesis
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   418
        by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   419
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   420
  qed
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
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   423
instantiation nat :: semiring_bits
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   424
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   425
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   426
definition bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   427
  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
   428
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   429
instance
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   430
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   431
  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
   432
    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
   433
    for P and n :: nat
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   434
  proof (induction n rule: nat_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   435
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   436
    from stable [of 0] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   437
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   438
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   439
    case (even n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   440
    with rec [of n False] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   441
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   442
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   443
    case (odd n)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   444
    with rec [of n True] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   445
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   446
  qed
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   447
qed (auto simp: div_mult2_eq bit_nat_def)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   448
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   449
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   450
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   451
lemma possible_bit_nat [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   452
  \<open>possible_bit TYPE(nat) n\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   453
  by (simp add: possible_bit_def)
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   454
79069
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
   455
lemma bit_Suc_0_iff [bit_simps]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
   456
  \<open>bit (Suc 0) n \<longleftrightarrow> n = 0\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
   457
  using bit_1_iff [of n, where ?'a = nat] by simp
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
   458
74497
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   459
lemma not_bit_Suc_0_Suc [simp]:
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   460
  \<open>\<not> bit (Suc 0) (Suc n)\<close>
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   461
  by (simp add: bit_Suc)
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   462
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   463
lemma not_bit_Suc_0_numeral [simp]:
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   464
  \<open>\<not> bit (Suc 0) (numeral n)\<close>
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   465
  by (simp add: numeral_eq_Suc)
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
   466
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   467
context semiring_bits
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   468
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   469
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   470
lemma bit_of_nat_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   471
  \<open>bit (of_nat m) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit m n\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   472
proof (cases \<open>possible_bit TYPE('a) n\<close>)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   473
  case False
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   474
  then show ?thesis
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   475
    by (simp add: impossible_bit)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   476
next
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   477
  case True
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   478
  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
   479
  proof (induction m arbitrary: n rule: nat_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   480
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   481
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   482
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   483
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   484
    case (even m)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   485
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   486
      by (cases n)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   487
        (auto simp: bit_double_iff Bit_Operations.bit_double_iff possible_bit_def bit_0 dest: mult_not_zero)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   488
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   489
    case (odd m)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   490
    then show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   491
      by (cases n)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   492
        (auto simp: bit_double_iff even_bit_succ_iff possible_bit_def
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
   493
          Bit_Operations.bit_Suc Bit_Operations.bit_0 dest: mult_not_zero)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   494
  qed
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   495
  with True show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   496
    by simp
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   497
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   498
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   499
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   500
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   501
lemma int_bit_induct [case_names zero minus even odd]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   502
  \<open>P k\<close> if zero_int: \<open>P 0\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   503
    and minus_int: \<open>P (- 1)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   504
    and even_int: \<open>\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   505
    and odd_int: \<open>\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))\<close> for k :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   506
proof (cases \<open>k \<ge> 0\<close>)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   507
  case True
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   508
  define n where \<open>n = nat k\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   509
  with True have \<open>k = int n\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   510
    by simp
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   511
  then show \<open>P k\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   512
  proof (induction n arbitrary: k rule: nat_bit_induct)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   513
    case zero
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   514
    then show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   515
      by (simp add: zero_int)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   516
  next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   517
    case (even n)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   518
    have \<open>P (int n * 2)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   519
      by (rule even_int) (use even in simp_all)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   520
    with even show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   521
      by (simp add: ac_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   522
  next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   523
    case (odd n)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   524
    have \<open>P (1 + (int n * 2))\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   525
      by (rule odd_int) (use odd in simp_all)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   526
    with odd show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   527
      by (simp add: ac_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   528
  qed
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   529
next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   530
  case False
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   531
  define n where \<open>n = nat (- k - 1)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   532
  with False have \<open>k = - int n - 1\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   533
    by simp
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   534
  then show \<open>P k\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   535
  proof (induction n arbitrary: k rule: nat_bit_induct)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   536
    case zero
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   537
    then show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   538
      by (simp add: minus_int)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   539
  next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   540
    case (even n)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   541
    have \<open>P (1 + (- int (Suc n) * 2))\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   542
      by (rule odd_int) (use even in \<open>simp_all add: algebra_simps\<close>)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   543
    also have \<open>\<dots> = - int (2 * n) - 1\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   544
      by (simp add: algebra_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   545
    finally show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   546
      using even.prems by simp
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   547
  next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   548
    case (odd n)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   549
    have \<open>P (- int (Suc n) * 2)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   550
      by (rule even_int) (use odd in \<open>simp_all add: algebra_simps\<close>)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   551
    also have \<open>\<dots> = - int (Suc (2 * n)) - 1\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   552
      by (simp add: algebra_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   553
    finally show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   554
      using odd.prems by simp
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   555
  qed
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   556
qed
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   557
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   558
instantiation int :: semiring_bits
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   559
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   560
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   561
definition bit_int :: \<open>int \<Rightarrow> nat \<Rightarrow> bool\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   562
  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
   563
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   564
instance
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   565
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   566
  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
   567
    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
   568
    for P and k :: int
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   569
  proof (induction k rule: int_bit_induct)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   570
    case zero
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   571
    from stable [of 0] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   572
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   573
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   574
    case minus
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   575
    from stable [of \<open>- 1\<close>] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   576
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   577
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   578
    case (even k)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   579
    with rec [of k False] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   580
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   581
  next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   582
    case (odd k)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   583
    with rec [of k True] show ?case
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   584
      by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   585
  qed
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   586
qed (auto simp: zdiv_zmult2_eq bit_int_def)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   587
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   588
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   589
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   590
lemma possible_bit_int [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   591
  \<open>possible_bit TYPE(int) n\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   592
  by (simp add: possible_bit_def)
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
   593
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   594
lemma bit_nat_iff [bit_simps]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   595
  \<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
   596
proof (cases \<open>k \<ge> 0\<close>)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   597
  case True
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   598
  moreover define m where \<open>m = nat k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   599
  ultimately have \<open>k = int m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   600
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   601
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   602
    by (simp add: bit_simps)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   603
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   604
  case False
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   605
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   606
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   607
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   608
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   609
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   610
subsection \<open>Bit operations\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   611
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   612
class semiring_bit_operations = semiring_bits +
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   613
  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
   614
    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
   615
    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
   616
    and mask :: \<open>nat \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   617
    and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   618
    and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   619
    and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   620
    and push_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   621
    and drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   622
    and take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   623
  assumes and_rec: \<open>a AND b = of_bool (odd a \<and> odd b) + 2 * ((a div 2) AND (b div 2))\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   624
    and or_rec: \<open>a OR b = of_bool (odd a \<or> odd b) + 2 * ((a div 2) OR (b div 2))\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   625
    and xor_rec: \<open>a XOR b = of_bool (odd a \<noteq> odd b) + 2 * ((a div 2) XOR (b div 2))\<close>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   626
    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
   627
    and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close>
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
   628
    and unset_bit_eq_or_xor: \<open>unset_bit n a = (a OR push_bit n 1) XOR push_bit n 1\<close>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   629
    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
   630
    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
   631
    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
   632
    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
   633
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   634
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   635
text \<open>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   636
  We want the bitwise operations to bind slightly weaker
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   637
  than \<open>+\<close> and \<open>-\<close>.
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   638
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   639
  Logically, \<^const>\<open>push_bit\<close>,
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   640
  \<^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
   641
  as separate operations makes proofs easier, otherwise proof automation
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   642
  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
   643
  algebraic relationships between those operations.
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   644
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
   645
  For the sake of code generation operations
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   646
  are specified as definitional class operations,
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   647
  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
   648
  differently wrt. code generation.
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   649
\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   650
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   651
lemma bit_iff_odd_drop_bit:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   652
  \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   653
  by (simp add: bit_iff_odd drop_bit_eq_div)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   654
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   655
lemma even_drop_bit_iff_not_bit:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   656
  \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   657
  by (simp add: bit_iff_odd_drop_bit)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   658
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   659
lemma bit_and_iff [bit_simps]:
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   660
  \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   661
proof (induction n arbitrary: a b)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   662
  case 0
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   663
  show ?case
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   664
    by (simp add: bit_0 and_rec [of a b] even_bit_succ_iff)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   665
next
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   666
  case (Suc n)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   667
  from Suc [of \<open>a div 2\<close> \<open>b div 2\<close>]
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   668
  show ?case
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   669
    by (simp add: and_rec [of a b] bit_Suc)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   670
      (auto simp flip: bit_Suc simp add: bit_double_iff dest: bit_imp_possible_bit)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   671
qed
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   672
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   673
lemma bit_or_iff [bit_simps]:
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   674
  \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   675
proof (induction n arbitrary: a b)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   676
  case 0
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   677
  show ?case
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   678
    by (simp add: bit_0 or_rec [of a b] even_bit_succ_iff)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   679
next
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   680
  case (Suc n)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   681
  from Suc [of \<open>a div 2\<close> \<open>b div 2\<close>]
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   682
  show ?case
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   683
    by (simp add: or_rec [of a b] bit_Suc)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   684
      (auto simp flip: bit_Suc simp add: bit_double_iff dest: bit_imp_possible_bit)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   685
qed
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   686
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   687
lemma bit_xor_iff [bit_simps]:
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   688
  \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   689
proof (induction n arbitrary: a b)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   690
  case 0
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   691
  show ?case
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   692
    by (simp add: bit_0 xor_rec [of a b] even_bit_succ_iff)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   693
next
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   694
  case (Suc n)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   695
  from Suc [of \<open>a div 2\<close> \<open>b div 2\<close>]
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   696
  show ?case
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   697
    by (simp add: xor_rec [of a b] bit_Suc)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   698
      (auto simp flip: bit_Suc simp add: bit_double_iff dest: bit_imp_possible_bit)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   699
qed
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
   700
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   701
sublocale "and": semilattice \<open>(AND)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   702
  by standard (auto simp: bit_eq_iff bit_and_iff)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   703
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   704
sublocale or: semilattice_neutr \<open>(OR)\<close> 0
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   705
  by standard (auto simp: bit_eq_iff bit_or_iff)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   706
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   707
sublocale xor: comm_monoid \<open>(XOR)\<close> 0
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   708
  by standard (auto simp: bit_eq_iff bit_xor_iff)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   709
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   710
lemma even_and_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   711
  \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   712
  using bit_and_iff [of a b 0] by (auto simp: bit_0)
74108
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
lemma even_or_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   715
  \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   716
  using bit_or_iff [of a b 0] by (auto simp: bit_0)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   717
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   718
lemma even_xor_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   719
  \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   720
  using bit_xor_iff [of a b 0] by (auto simp: bit_0)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   721
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   722
lemma zero_and_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   723
  \<open>0 AND a = 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   724
  by (simp add: bit_eq_iff bit_and_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   725
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   726
lemma and_zero_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   727
  \<open>a AND 0 = 0\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   728
  by (simp add: bit_eq_iff bit_and_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   729
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   730
lemma one_and_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   731
  \<open>1 AND a = a mod 2\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   732
  by (simp add: bit_eq_iff bit_and_iff) (auto simp: bit_1_iff bit_0)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   733
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   734
lemma and_one_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   735
  \<open>a AND 1 = a mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   736
  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
   737
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   738
lemma one_or_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   739
  \<open>1 OR a = a + of_bool (even a)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
   740
  by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   741
    (auto simp: bit_1_iff bit_0)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   742
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   743
lemma or_one_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   744
  \<open>a OR 1 = a + of_bool (even a)\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   745
  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
   746
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   747
lemma one_xor_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   748
  \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
   749
  by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   750
    (auto simp: bit_1_iff odd_bit_iff_bit_pred bit_0 elim: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   751
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   752
lemma xor_one_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
   753
  \<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
   754
  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
   755
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
   756
lemma xor_self_eq [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
   757
  \<open>a XOR a = 0\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
   758
  by (rule bit_eqI) (simp add: bit_simps)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
   759
79588
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   760
lemma mask_0 [simp]:
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   761
  \<open>mask 0 = 0\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   762
  by (simp add: mask_eq_exp_minus_1)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   763
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   764
lemma inc_mask_eq_exp:
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   765
  \<open>mask n + 1 = 2 ^ n\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   766
proof (induction n)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   767
  case 0
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   768
  then show ?case
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   769
    by simp
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   770
next
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   771
  case (Suc n)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   772
  from Suc.IH [symmetric] have \<open>2 ^ Suc n = 2 * mask n + 2\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   773
    by (simp add: algebra_simps)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   774
  also have \<open>\<dots> = 2 * mask n + 1 + 1\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   775
    by (simp add: add.assoc)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   776
  finally have *: \<open>2 ^ Suc n = 2 * mask n + 1 + 1\<close> .
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   777
  then show ?case
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   778
    by (simp add: mask_eq_exp_minus_1)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   779
qed
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   780
82524
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   781
lemma mask_eq_iff_eq_exp:
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   782
  \<open>mask n = a \<longleftrightarrow> a + 1 = 2 ^ n\<close>
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   783
  by (auto simp flip: inc_mask_eq_exp)
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   784
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   785
lemma eq_mask_iff_eq_exp:
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   786
  \<open>a = mask n \<longleftrightarrow> a + 1 = 2 ^ n\<close>
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   787
  by (auto simp flip: inc_mask_eq_exp)
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
   788
79588
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   789
lemma mask_Suc_double:
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   790
  \<open>mask (Suc n) = 1 OR 2 * mask n\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   791
proof -
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   792
  have \<open>mask (Suc n) + 1 = (mask n + 1) + (mask n + 1)\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   793
    by (simp add: inc_mask_eq_exp mult_2)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   794
  also have \<open>\<dots> = (1 OR 2 * mask n) + 1\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   795
    by (simp add: one_or_eq mult_2_right algebra_simps)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   796
  finally show ?thesis
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   797
    by simp
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   798
qed
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   799
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   800
lemma bit_mask_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   801
  \<open>bit (mask m) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n < m\<close>
79588
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   802
proof (cases \<open>possible_bit TYPE('a) n\<close>)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   803
  case False
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   804
  then show ?thesis
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   805
    by (simp add: impossible_bit)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   806
next
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   807
  case True
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   808
  then have \<open>bit (mask m) n \<longleftrightarrow> n < m\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   809
  proof (induction m arbitrary: n)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   810
    case 0
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   811
    then show ?case
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   812
      by (simp add: bit_iff_odd)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   813
  next
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   814
    case (Suc m)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   815
    show ?case
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   816
    proof (cases n)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   817
      case 0
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   818
      then show ?thesis
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   819
        by (simp add: bit_0 mask_Suc_double even_or_iff)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   820
    next
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   821
      case (Suc n)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   822
      with Suc.prems have \<open>possible_bit TYPE('a) n\<close>
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   823
        using possible_bit_less_imp by auto
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   824
      with Suc.IH [of n] have \<open>bit (mask m) n \<longleftrightarrow> n < m\<close> .
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   825
      with Suc.prems show ?thesis
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   826
        by (simp add: Suc mask_Suc_double bit_simps)
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   827
    qed
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   828
  qed
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   829
  with True show ?thesis
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   830
    by simp
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   831
qed
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   832
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   833
lemma even_mask_iff:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   834
  \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   835
  using bit_mask_iff [of n 0] by (auto simp: bit_0)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   836
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   837
lemma mask_Suc_0 [simp]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   838
  \<open>mask (Suc 0) = 1\<close>
79588
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
   839
  by (simp add: mask_Suc_double)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   840
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   841
lemma mask_Suc_exp:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   842
  \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   843
  by (auto simp: bit_eq_iff bit_simps)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   844
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   845
lemma mask_numeral:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   846
  \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   847
  by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   848
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   849
lemma push_bit_0_id [simp]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   850
  \<open>push_bit 0 = id\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   851
  by (simp add: fun_eq_iff push_bit_eq_mult)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   852
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   853
lemma push_bit_Suc [simp]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   854
  \<open>push_bit (Suc n) a = push_bit n (a * 2)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   855
  by (simp add: push_bit_eq_mult ac_simps)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   856
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   857
lemma push_bit_double:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   858
  \<open>push_bit n (a * 2) = push_bit n a * 2\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   859
  by (simp add: push_bit_eq_mult ac_simps)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   860
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   861
lemma bit_push_bit_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   862
  \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> possible_bit TYPE('a) n \<and> bit a (n - m)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   863
proof (induction n arbitrary: m)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   864
  case 0
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   865
  then show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   866
    by (auto simp: bit_0 push_bit_eq_mult)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   867
next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   868
  case (Suc n)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   869
  show ?case
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   870
  proof (cases m)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   871
    case 0
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   872
    then show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   873
      by (auto simp: bit_imp_possible_bit)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   874
  next
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   875
    case (Suc m')
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   876
    with Suc.prems Suc.IH [of m'] show ?thesis
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   877
      apply (simp add: push_bit_double)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   878
      apply (auto simp: possible_bit_less_imp bit_simps mult.commute [of _ 2])
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   879
      done
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   880
  qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   881
qed
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   882
79590
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
   883
lemma funpow_double_eq_push_bit:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
   884
  \<open>times 2 ^^ n = push_bit n\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
   885
  by (induction n) (simp_all add: fun_eq_iff push_bit_double ac_simps)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
   886
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   887
lemma div_push_bit_of_1_eq_drop_bit:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   888
  \<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
   889
  by (simp add: push_bit_eq_mult drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   890
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   891
lemma bits_ident:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   892
  \<open>push_bit n (drop_bit n a) + take_bit n a = a\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   893
  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
   894
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   895
lemma push_bit_push_bit [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   896
  \<open>push_bit m (push_bit n a) = push_bit (m + n) a\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   897
  by (simp add: push_bit_eq_mult power_add ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   898
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   899
lemma push_bit_of_0 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   900
  \<open>push_bit n 0 = 0\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   901
  by (simp add: push_bit_eq_mult)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   902
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
   903
lemma push_bit_of_1 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   904
  \<open>push_bit n 1 = 2 ^ n\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   905
  by (simp add: push_bit_eq_mult)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   906
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   907
lemma push_bit_add:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
   908
  \<open>push_bit n (a + b) = push_bit n a + push_bit n b\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   909
  by (simp add: push_bit_eq_mult algebra_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   910
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   911
lemma push_bit_numeral [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   912
  \<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
   913
  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
   914
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   915
lemma bit_drop_bit_eq [bit_simps]:
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   916
  \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   917
  by rule (simp add: drop_bit_eq_div bit_iff_odd div_exp_eq)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   918
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   919
lemma disjunctive_xor_eq_or:
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   920
  \<open>a XOR b = a OR b\<close> if \<open>a AND b = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   921
  using that by (auto simp: bit_eq_iff bit_simps)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   922
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   923
lemma disjunctive_add_eq_or:
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   924
  \<open>a + b = a OR b\<close> if \<open>a AND b = 0\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   925
proof (rule bit_eqI)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   926
  fix n
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   927
  assume \<open>possible_bit TYPE('a) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   928
  moreover from that have \<open>\<And>n. \<not> bit (a AND b) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   929
    by simp
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   930
  then have \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   931
    by (simp add: bit_simps)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   932
  ultimately show \<open>bit (a + b) n \<longleftrightarrow> bit (a OR b) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   933
  proof (induction n arbitrary: a b)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   934
    case 0
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   935
    from "0"(2)[of 0] show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   936
      by (auto simp: even_or_iff bit_0)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   937
  next
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   938
    case (Suc n)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   939
    from Suc.prems(2) [of 0] have even: \<open>even a \<or> even b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   940
      by (auto simp: bit_0)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   941
    have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   942
      using Suc.prems(2) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   943
    from Suc.prems have \<open>possible_bit TYPE('a) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   944
      using possible_bit_less_imp by force
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   945
    with \<open>\<And>n. \<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> Suc.IH [of \<open>a div 2\<close> \<open>b div 2\<close>]
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   946
    have IH: \<open>bit (a div 2 + b div 2) n \<longleftrightarrow> bit (a div 2 OR b div 2) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   947
      by (simp add: bit_Suc)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   948
    have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   949
      using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   950
    also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
   951
      using even by (auto simp: algebra_simps mod2_eq_if)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   952
    finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   953
      using \<open>possible_bit TYPE('a) (Suc n)\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff possible_bit_def)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   954
    also have \<open>\<dots> \<longleftrightarrow> bit (a div 2 OR b div 2) n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   955
      by (rule IH)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   956
    finally show ?case
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   957
      by (simp add: bit_simps flip: bit_Suc)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   958
  qed
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   959
qed
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   960
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   961
lemma disjunctive_add_eq_xor:
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   962
  \<open>a + b = a XOR b\<close> if \<open>a AND b = 0\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   963
  using that by (simp add: disjunctive_add_eq_or disjunctive_xor_eq_or)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   964
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   965
lemma take_bit_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   966
  "take_bit 0 a = 0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   967
  by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
   968
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   969
lemma bit_take_bit_iff [bit_simps]:
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
   970
  \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close>
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   971
proof -
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   972
  have \<open>push_bit m (drop_bit m a) AND take_bit m a = 0\<close> (is \<open>?lhs = _\<close>)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   973
  proof (rule bit_eqI)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   974
    fix n
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   975
    show \<open>bit ?lhs n \<longleftrightarrow> bit 0 n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   976
    proof (cases \<open>m \<le> n\<close>)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   977
      case False
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   978
      then show ?thesis
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   979
        by (simp add: bit_simps)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   980
    next
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   981
      case True
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   982
      moreover define q where \<open>q = n - m\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   983
      ultimately have \<open>n = m + q\<close> by simp
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   984
      moreover have \<open>\<not> bit (take_bit m a) (m + q)\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   985
        by (simp add: take_bit_eq_mod bit_iff_odd flip: div_exp_eq)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   986
      ultimately show ?thesis
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   987
        by (simp add: bit_simps)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   988
    qed
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   989
  qed
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   990
  then have \<open>push_bit m (drop_bit m a) XOR take_bit m a = push_bit m (drop_bit m a) + take_bit m a\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   991
    by (simp add: disjunctive_add_eq_xor)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   992
  also have \<open>\<dots> = a\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   993
    by (simp add: bits_ident)
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   994
  finally have \<open>bit (push_bit m (drop_bit m a) XOR take_bit m a) n \<longleftrightarrow> bit a n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   995
    by simp
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   996
  also have \<open>\<dots> \<longleftrightarrow> (m \<le> n \<or> n < m) \<and> bit a n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   997
    by auto
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   998
  also have \<open>\<dots> \<longleftrightarrow> m \<le> n \<and> bit a n \<or> n < m \<and> bit a n\<close>
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
   999
    by auto
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
  1000
  also have \<open>m \<le> n \<and> bit a n \<longleftrightarrow> bit (push_bit m (drop_bit m a)) n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1001
    by (auto simp: bit_simps bit_imp_possible_bit)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
  1002
  finally show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1003
    by (auto simp: bit_simps)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
  1004
qed
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1005
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1006
lemma take_bit_Suc:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1007
  \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\<close> (is \<open>?lhs = ?rhs\<close>)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1008
proof (rule bit_eqI)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1009
  fix m
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1010
  assume \<open>possible_bit TYPE('a) m\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1011
  then show \<open>bit ?lhs m \<longleftrightarrow> bit ?rhs m\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1012
    apply (cases a rule: parity_cases; cases m)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1013
       apply (simp_all add: bit_simps even_bit_succ_iff mult.commute [of _ 2] add.commute [of _ 1] flip: bit_Suc)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1014
     apply (simp_all add: bit_0)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1015
    done
74101
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
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1018
lemma take_bit_rec:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1019
  \<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
  1020
  by (cases n) (simp_all add: take_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1021
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1022
lemma take_bit_Suc_0 [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1023
  \<open>take_bit (Suc 0) a = a mod 2\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1024
  by (simp add: take_bit_eq_mod)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1025
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1026
lemma take_bit_of_0 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1027
  \<open>take_bit n 0 = 0\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1028
  by (rule bit_eqI) (simp add: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1029
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1030
lemma take_bit_of_1 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1031
  \<open>take_bit n 1 = of_bool (n > 0)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1032
  by (cases n) (simp_all add: take_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1033
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1034
lemma drop_bit_of_0 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1035
  \<open>drop_bit n 0 = 0\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1036
  by (rule bit_eqI) (simp add: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1037
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1038
lemma drop_bit_of_1 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1039
  \<open>drop_bit n 1 = of_bool (n = 0)\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1040
  by (rule bit_eqI) (simp add: bit_simps ac_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1041
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1042
lemma drop_bit_0 [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1043
  \<open>drop_bit 0 = id\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1044
  by (simp add: fun_eq_iff drop_bit_eq_div)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1045
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1046
lemma drop_bit_Suc:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1047
  \<open>drop_bit (Suc n) a = drop_bit n (a div 2)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1048
  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
  1049
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1050
lemma drop_bit_rec:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1051
  \<open>drop_bit n a = (if n = 0 then a else drop_bit (n - 1) (a div 2))\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1052
  by (cases n) (simp_all add: drop_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1053
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1054
lemma drop_bit_half:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1055
  \<open>drop_bit n (a div 2) = drop_bit n a div 2\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1056
  by (induction n arbitrary: a) (simp_all add: drop_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1057
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1058
lemma drop_bit_of_bool [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1059
  \<open>drop_bit n (of_bool b) = of_bool (n = 0 \<and> b)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1060
  by (cases n) simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1061
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1062
lemma even_take_bit_eq [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1063
  \<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
  1064
  by (simp add: take_bit_rec [of n a])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1065
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1066
lemma take_bit_take_bit [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1067
  \<open>take_bit m (take_bit n a) = take_bit (min m n) a\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1068
  by (rule bit_eqI) (simp add: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1069
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1070
lemma drop_bit_drop_bit [simp]:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1071
  \<open>drop_bit m (drop_bit n a) = drop_bit (m + n) a\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1072
  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
  1073
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1074
lemma push_bit_take_bit:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1075
  \<open>push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1076
  by (rule bit_eqI) (auto simp: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1077
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1078
lemma take_bit_push_bit:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1079
  \<open>take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1080
  by (rule bit_eqI) (auto simp: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1081
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1082
lemma take_bit_drop_bit:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1083
  \<open>take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1084
  by (rule bit_eqI) (auto simp: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1085
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1086
lemma drop_bit_take_bit:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1087
  \<open>drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1088
  by (rule bit_eqI) (auto simp: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1089
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1090
lemma even_push_bit_iff [simp]:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1091
  \<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
  1092
  by (simp add: push_bit_eq_mult) auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1093
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1094
lemma stable_imp_drop_bit_eq:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1095
  \<open>drop_bit n a = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1096
  if \<open>a div 2 = a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1097
  by (induction n) (simp_all add: that drop_bit_Suc)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1098
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1099
lemma stable_imp_take_bit_eq:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1100
  \<open>take_bit n a = (if even a then 0 else mask n)\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1101
    if \<open>a div 2 = a\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1102
  by (rule bit_eqI) (use that in \<open>simp add: bit_simps stable_imp_bit_iff_odd\<close>)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1103
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1104
lemma exp_dvdE:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1105
  assumes \<open>2 ^ n dvd a\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1106
  obtains b where \<open>a = push_bit n b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1107
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1108
  from assms obtain b where \<open>a = 2 ^ n * b\<close> ..
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1109
  then have \<open>a = push_bit n b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1110
    by (simp add: push_bit_eq_mult ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1111
  with that show thesis .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1112
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1113
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1114
lemma take_bit_eq_0_iff:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1115
  \<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
  1116
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1117
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1118
  then show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1119
    by (simp add: take_bit_eq_mod mod_0_imp_dvd)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1120
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1121
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1122
  then obtain b where \<open>a = push_bit n b\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1123
    by (rule exp_dvdE)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1124
  then show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1125
    by (simp add: take_bit_push_bit)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1126
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1127
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1128
lemma take_bit_tightened:
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  1129
  \<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>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1130
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1131
  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
  1132
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1133
  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
  1134
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1135
  with that show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1136
    by (simp add: min_def)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1137
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1138
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1139
lemma take_bit_eq_self_iff_drop_bit_eq_0:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1140
  \<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
  1141
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1142
  assume ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1143
  show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1144
  proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1145
    fix m
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1146
    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
  1147
    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
  1148
      unfolding bit_simps
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  1149
      by (simp add: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1150
    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
  1151
      by (simp add: bit_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1152
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1153
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1154
  assume ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1155
  show ?P
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1156
  proof (rule bit_eqI)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1157
    fix m
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1158
    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
  1159
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1160
    then have \<open> \<not> bit a (n + (m - n))\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1161
      by (simp add: bit_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1162
    then show \<open>bit (take_bit n a) m \<longleftrightarrow> bit a m\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1163
      by (cases \<open>m < n\<close>) (auto simp: bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1164
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1165
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1166
82289
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1167
lemma impossible_bit_imp_take_bit_eq_self:
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1168
  \<open>take_bit n a = a\<close> if \<open>\<not> possible_bit TYPE('a) n\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1169
proof -
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1170
  have \<open>drop_bit n a = 0\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1171
  proof (rule bit_eqI)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1172
    fix m
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1173
    show \<open>bit (drop_bit n a) m \<longleftrightarrow> bit 0 m\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1174
      using possible_bit_less_imp [of \<open>n + m\<close> n] that
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1175
      by (auto simp add: bit_simps dest: bit_imp_possible_bit)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1176
  qed
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1177
  then show ?thesis
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1178
    by (simp add: take_bit_eq_self_iff_drop_bit_eq_0)
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1179
qed
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1180
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1181
lemma drop_bit_exp_eq:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1182
  \<open>drop_bit m (2 ^ n) = of_bool (m \<le> n \<and> possible_bit TYPE('a) n) * 2 ^ (n - m)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1183
  by (auto simp: bit_eq_iff bit_simps)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1184
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1185
lemma take_bit_and [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1186
  \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1187
  by (auto simp: bit_eq_iff bit_simps)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1188
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1189
lemma take_bit_or [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1190
  \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1191
  by (auto simp: bit_eq_iff bit_simps)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1192
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1193
lemma take_bit_xor [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1194
  \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1195
  by (auto simp: bit_eq_iff bit_simps)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1196
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1197
lemma push_bit_and [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1198
  \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1199
  by (auto simp: bit_eq_iff bit_simps)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1200
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1201
lemma push_bit_or [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1202
  \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1203
  by (auto simp: bit_eq_iff bit_simps)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1204
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1205
lemma push_bit_xor [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1206
  \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1207
  by (auto simp: bit_eq_iff bit_simps)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1208
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1209
lemma drop_bit_and [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1210
  \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1211
  by (auto simp: bit_eq_iff bit_simps)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1212
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1213
lemma drop_bit_or [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1214
  \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1215
  by (auto simp: bit_eq_iff bit_simps)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1216
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1217
lemma drop_bit_xor [simp]:
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1218
  \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1219
  by (auto simp: bit_eq_iff bit_simps)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1220
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1221
lemma take_bit_of_mask [simp]:
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1222
  \<open>take_bit m (mask n) = mask (min m n)\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1223
  by (rule bit_eqI) (simp add: bit_simps)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  1224
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
  1225
lemma take_bit_eq_mask:
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1226
  \<open>take_bit n a = a AND mask n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1227
  by (auto simp: bit_eq_iff bit_simps)
71823
214b48a1937b explicit mask operation for bits
haftmann
parents: 71822
diff changeset
  1228
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1229
lemma or_eq_0_iff:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1230
  \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1231
  by (auto simp: 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
  1232
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1233
lemma bit_iff_and_drop_bit_eq_1:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1234
  \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1235
  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
  1236
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1237
lemma bit_iff_and_push_bit_not_eq_0:
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1238
  \<open>bit a n \<longleftrightarrow> a AND push_bit n 1 \<noteq> 0\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  1239
  by (cases \<open>possible_bit TYPE('a) n\<close>) (simp_all add: bit_eq_iff bit_simps impossible_bit)
72508
c89d8e8bd8c7 factored out theory Traditional_Syntax
haftmann
parents: 72488
diff changeset
  1240
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1241
lemma bit_set_bit_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1242
  \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> possible_bit TYPE('a) n)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1243
  by (auto simp: set_bit_eq_or bit_or_iff bit_exp_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1244
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1245
lemma even_set_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1246
  \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1247
  using bit_set_bit_iff [of m a 0] by (auto simp: bit_0)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1248
79031
4596a14d9a95 slightly more elementary characterization of unset_bit
haftmann
parents: 79030
diff changeset
  1249
lemma bit_unset_bit_iff [bit_simps]:
4596a14d9a95 slightly more elementary characterization of unset_bit
haftmann
parents: 79030
diff changeset
  1250
  \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1251
  by (auto simp: unset_bit_eq_or_xor bit_simps dest: bit_imp_possible_bit)
79031
4596a14d9a95 slightly more elementary characterization of unset_bit
haftmann
parents: 79030
diff changeset
  1252
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1253
lemma even_unset_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1254
  \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1255
  using bit_unset_bit_iff [of m a 0] by (auto simp: bit_0)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1256
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1257
lemma bit_flip_bit_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1258
  \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> possible_bit TYPE('a) n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1259
  by (auto simp: bit_eq_iff bit_simps flip_bit_eq_xor bit_imp_possible_bit)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1260
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1261
lemma even_flip_bit_iff:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1262
  \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  1263
  using bit_flip_bit_iff [of m a 0] by (auto simp: possible_bit_def  bit_0)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1264
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1265
lemma and_exp_eq_0_iff_not_bit:
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1266
  \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1267
  using bit_imp_possible_bit[of a n]
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1268
  by (auto simp: bit_eq_iff bit_simps)
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1269
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1270
lemma bit_sum_mult_2_cases:
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1271
  assumes a: \<open>\<forall>j. \<not> bit a (Suc j)\<close>
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1272
  shows \<open>bit (a + 2 * b) n = (if n = 0 then odd a else bit (2 * b) n)\<close>
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1273
proof -
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1274
  from a have \<open>n = 0\<close> if \<open>bit a n\<close> for n using that
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1275
    by (cases n) simp_all
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1276
  then have \<open>a = 0 \<or> a = 1\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1277
    by (auto simp: bit_eq_iff bit_1_iff)
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1278
  then show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1279
    by (cases n) (auto simp: bit_0 bit_double_iff even_bit_succ_iff)
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1280
qed
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1281
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1282
lemma set_bit_0:
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1283
  \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1284
  by (auto simp: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc)
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1285
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1286
lemma set_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1287
  \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1288
  by (auto simp: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1289
    elim: possible_bit_less_imp)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1290
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1291
lemma unset_bit_0:
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1292
  \<open>unset_bit 0 a = 2 * (a div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1293
  by (auto simp: bit_eq_iff bit_simps simp flip: bit_Suc)
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1294
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1295
lemma unset_bit_Suc:
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1296
  \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1297
  by (auto simp: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc)
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1298
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1299
lemma flip_bit_0:
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1300
  \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1301
  by (auto simp: bit_eq_iff bit_simps even_bit_succ_iff bit_0 simp flip: bit_Suc)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1302
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1303
lemma flip_bit_Suc:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1304
  \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1305
  by (auto simp: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1306
    elim: possible_bit_less_imp)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1307
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1308
lemma flip_bit_eq_if:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1309
  \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1310
  by (rule bit_eqI) (auto simp: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1311
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1312
lemma take_bit_set_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1313
  \<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>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1314
  by (rule bit_eqI) (auto simp: bit_take_bit_iff bit_set_bit_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1315
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1316
lemma take_bit_unset_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1317
  \<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>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1318
  by (rule bit_eqI) (auto simp: bit_take_bit_iff bit_unset_bit_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1319
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1320
lemma take_bit_flip_bit_eq:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1321
  \<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>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1322
  by (rule bit_eqI) (auto simp: bit_take_bit_iff bit_flip_bit_iff)
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1323
74497
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1324
lemma push_bit_Suc_numeral [simp]:
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1325
  \<open>push_bit (Suc n) (numeral k) = push_bit n (numeral (Num.Bit0 k))\<close>
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1326
  by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0)
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1327
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1328
lemma mask_eq_0_iff [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1329
  \<open>mask n = 0 \<longleftrightarrow> n = 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1330
  by (cases n) (simp_all add: mask_Suc_double or_eq_0_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1331
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1332
lemma bit_horner_sum_bit_iff [bit_simps]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1333
  \<open>bit (horner_sum of_bool 2 bs) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n < length bs \<and> bs ! n\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1334
proof (induction bs arbitrary: n)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1335
  case Nil
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1336
  then show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1337
    by simp
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1338
next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1339
  case (Cons b bs)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1340
  show ?case
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1341
  proof (cases n)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1342
    case 0
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1343
    then show ?thesis
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1344
      by (simp add: bit_0)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1345
  next
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1346
    case (Suc m)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1347
    with bit_rec [of _ n] Cons.prems Cons.IH [of m]
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1348
    show ?thesis
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1349
      by (simp add: bit_simps)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1350
        (auto simp: possible_bit_less_imp bit_simps simp flip: bit_Suc)
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1351
  qed
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1352
qed
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1353
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1354
lemma horner_sum_bit_eq_take_bit:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1355
  \<open>horner_sum of_bool 2 (map (bit a) [0..<n]) = take_bit n a\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1356
  by (rule bit_eqI) (auto simp: bit_simps)
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1357
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1358
lemma take_bit_horner_sum_bit_eq:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1359
  \<open>take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1360
  by (auto simp: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff)
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1361
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1362
lemma take_bit_sum:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1363
  \<open>take_bit n a = (\<Sum>k = 0..<n. push_bit k (of_bool (bit a k)))\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1364
  by (simp flip: horner_sum_bit_eq_take_bit add: horner_sum_eq_sum push_bit_eq_mult)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1365
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1366
lemma set_bit_eq:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1367
  \<open>set_bit n a = a + of_bool (\<not> bit a n) * 2 ^ n\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1368
proof -
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1369
  have \<open>a AND of_bool (\<not> bit a n) * 2 ^ n = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1370
    by (auto simp: bit_eq_iff bit_simps)
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1371
  then show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1372
    by (auto simp: bit_eq_iff bit_simps disjunctive_add_eq_or)
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1373
qed
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1374
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1375
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1376
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1377
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
  1378
  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1379
  assumes not_eq_complement: \<open>NOT a = - 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
  1380
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1381
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1382
text \<open>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1383
  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
  1384
  definitional class operation.  Note that \<^const>\<open>not\<close> has no
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1385
  sensible definition for unlimited but only positive bit strings
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1386
  (type \<^typ>\<open>nat\<close>).
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1387
\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1388
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1389
lemma bits_minus_1_mod_2_eq [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1390
  \<open>(- 1) mod 2 = 1\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1391
  by (simp add: mod_2_eq_odd)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1392
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1393
lemma minus_eq_not_plus_1:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1394
  \<open>- a = NOT a + 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1395
  using not_eq_complement [of a] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1396
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1397
lemma minus_eq_not_minus_1:
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1398
  \<open>- a = NOT (a - 1)\<close>
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1399
  using not_eq_complement [of \<open>a - 1\<close>] by simp (simp add: algebra_simps)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1400
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1401
lemma not_rec:
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1402
  \<open>NOT a = of_bool (even a) + 2 * NOT (a div 2)\<close>
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1403
  by (simp add: not_eq_complement algebra_simps mod_2_eq_odd flip: minus_mod_eq_mult_div)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1404
79590
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1405
lemma decr_eq_not_minus:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1406
  \<open>a - 1 = NOT (- a)\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1407
  using not_eq_complement [of \<open>- a\<close>] by simp
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1408
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1409
lemma even_not_iff [simp]:
79018
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1410
  \<open>even (NOT a) \<longleftrightarrow> odd a\<close>
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1411
  by (simp add: not_eq_complement)
79018
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1412
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1413
lemma bit_not_iff [bit_simps]:
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1414
  \<open>bit (NOT a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> \<not> bit a n\<close>
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1415
proof (cases \<open>possible_bit TYPE('a) n\<close>)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1416
  case False
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1417
  then show ?thesis
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  1418
    by (auto dest: bit_imp_possible_bit)
79018
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1419
next
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1420
  case True
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1421
  moreover have \<open>bit (NOT a) n \<longleftrightarrow> \<not> bit a n\<close>
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1422
  using \<open>possible_bit TYPE('a) n\<close> proof (induction n arbitrary: a)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1423
    case 0
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1424
    then show ?case
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1425
      by (simp add: bit_0)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1426
  next
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1427
    case (Suc n)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1428
    from Suc.prems Suc.IH [of \<open>a div 2\<close>]
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1429
    show ?case
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1430
      by (simp add: impossible_bit possible_bit_less_imp not_rec [of a] even_bit_succ_iff bit_double_iff flip: bit_Suc)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1431
  qed
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1432
  ultimately show ?thesis
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1433
    by simp
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1434
qed
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1435
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1436
lemma bit_not_exp_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1437
  \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<noteq> m\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1438
  by (auto simp: bit_not_iff bit_exp_iff)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1439
79018
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1440
lemma bit_minus_iff [bit_simps]:
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1441
  \<open>bit (- a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> \<not> bit (a - 1) n\<close>
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1442
  by (simp add: minus_eq_not_minus_1 bit_not_iff)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1443
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1444
lemma bit_minus_1_iff [simp]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1445
  \<open>bit (- 1) n \<longleftrightarrow> possible_bit TYPE('a) n\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1446
  by (simp add: bit_minus_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1447
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  1448
lemma bit_minus_exp_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1449
  \<open>bit (- (2 ^ m)) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<ge> m\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1450
  by (auto simp: bit_simps simp flip: mask_eq_exp_minus_1)
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1451
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1452
lemma bit_minus_2_iff [simp]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1453
  \<open>bit (- 2) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n > 0\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1454
  by (simp add: bit_minus_iff bit_1_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
  1455
79590
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1456
lemma bit_decr_iff:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1457
  \<open>bit (a - 1) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> \<not> bit (- a) n\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1458
  by (simp add: decr_eq_not_minus bit_not_iff)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  1459
79018
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1460
lemma bit_not_iff_eq:
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1461
  \<open>bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1462
  by (simp add: bit_simps possible_bit_def)
7449ff77029e base abstract specification of NOT on recursive equation rather than bit projection
haftmann
parents: 79017
diff changeset
  1463
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  1464
lemma not_one_eq [simp]:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1465
  \<open>NOT 1 = - 2\<close>
81641
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  1466
  by (rule bit_eqI, simp add: bit_simps)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1467
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1468
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1469
  by standard (rule bit_eqI, simp add: bit_and_iff)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1470
74123
7c5842b06114 clarified abstract and concrete boolean algebras
haftmann
parents: 74108
diff changeset
  1471
sublocale bit: abstract_boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1472
  by standard (auto simp: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
74123
7c5842b06114 clarified abstract and concrete boolean algebras
haftmann
parents: 74108
diff changeset
  1473
7c5842b06114 clarified abstract and concrete boolean algebras
haftmann
parents: 74108
diff changeset
  1474
sublocale bit: abstract_boolean_algebra_sym_diff \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close> \<open>(XOR)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1475
proof
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1476
  show \<open>\<And>x y. x XOR y = x AND NOT y OR NOT x AND y\<close>
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1477
    by (intro bit_eqI) (auto simp: bit_simps)
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1478
qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1479
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1480
lemma and_eq_not_not_or:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1481
  \<open>a AND b = NOT (NOT a OR NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1482
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1483
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1484
lemma or_eq_not_not_and:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1485
  \<open>a OR b = NOT (NOT a AND NOT b)\<close>
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1486
  by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1487
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1488
lemma not_add_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1489
  \<open>NOT (a + b) = NOT a - b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1490
  by (simp add: not_eq_complement algebra_simps)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1491
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1492
lemma not_diff_distrib:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1493
  \<open>NOT (a - b) = NOT a + b\<close>
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1494
  using not_add_distrib [of a \<open>- b\<close>] by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  1495
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1496
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
  1497
  \<open>a AND b = - 1 \<longleftrightarrow> a = - 1 \<and> b = - 1\<close>
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  1498
  by (auto simp: bit_eq_iff bit_simps)
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 72262
diff changeset
  1499
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1500
lemma disjunctive_and_not_eq_xor:
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1501
  \<open>a AND NOT b = a XOR b\<close> if \<open>NOT a AND b = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1502
  using that by (auto simp: bit_eq_iff bit_simps)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1503
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1504
lemma disjunctive_diff_eq_and_not:
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1505
  \<open>a - b = a AND NOT b\<close> if \<open>NOT a AND b = 0\<close>
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1506
proof -
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1507
  from that have \<open>NOT a + b = NOT a OR b\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1508
    by (rule disjunctive_add_eq_or)
72239
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1509
  then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1510
    by simp
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1511
  then show ?thesis
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1512
    by (simp add: not_add_distrib)
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1513
qed
12e94c2ff6c5 generalized
haftmann
parents: 72227
diff changeset
  1514
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1515
lemma disjunctive_diff_eq_xor:
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1516
  \<open>a AND NOT b = a XOR b\<close> if \<open>NOT a AND b = 0\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1517
  using that by (simp add: disjunctive_and_not_eq_xor disjunctive_diff_eq_and_not)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1518
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1519
lemma push_bit_minus:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1520
  \<open>push_bit n (- a) = - push_bit n a\<close>
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1521
  by (simp add: push_bit_eq_mult)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
  1522
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1523
lemma take_bit_not_take_bit:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
  1524
  \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1525
  by (auto simp: 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
  1526
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1527
lemma take_bit_not_iff:
73969
ca2a35c0fe6e operations for symbolic computation of bit operations
haftmann
parents: 73871
diff changeset
  1528
  \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1529
  by (auto simp: bit_eq_iff bit_simps)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
  1530
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1531
lemma take_bit_not_eq_mask_diff:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1532
  \<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
  1533
proof -
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1534
  have \<open>NOT (mask n) AND take_bit n a = 0\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1535
    by (simp add: bit_eq_iff bit_simps)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1536
  moreover have \<open>take_bit n (NOT a) = mask n AND NOT (take_bit n a)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1537
    by (auto simp: bit_eq_iff bit_simps)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1538
  ultimately show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1539
    by (simp add: disjunctive_diff_eq_and_not)
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1540
qed
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  1541
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1542
lemma mask_eq_take_bit_minus_one:
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1543
  \<open>mask n = take_bit n (- 1)\<close>
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1544
  by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1545
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1546
lemma take_bit_minus_one_eq_mask [simp]:
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1547
  \<open>take_bit n (- 1) = mask n\<close>
72079
8c355e2dd7db more consequent transferability
haftmann
parents: 72028
diff changeset
  1548
  by (simp add: mask_eq_take_bit_minus_one)
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1549
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1550
lemma minus_exp_eq_not_mask:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1551
  \<open>- (2 ^ n) = NOT (mask n)\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1552
  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
  1553
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1554
lemma push_bit_minus_one_eq_not_mask [simp]:
71922
2c6a5c709f22 more theorems
haftmann
parents: 71921
diff changeset
  1555
  \<open>push_bit n (- 1) = NOT (mask n)\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1556
  by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1557
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1558
lemma take_bit_not_mask_eq_0:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1559
  \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  1560
  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
  1561
82289
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1562
lemma mask_eq_minus_one_if_not_possible_bit:
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1563
  \<open>mask n = - 1\<close> if \<open>\<not> possible_bit TYPE('a) n\<close>
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1564
  using that mask_eq_take_bit_minus_one [of n] impossible_bit_imp_take_bit_eq_self [of n \<open>- 1\<close>]
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1565
  by simp
26fbbf43863b more lemmas
haftmann
parents: 81876
diff changeset
  1566
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1567
lemma unset_bit_eq_and_not:
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1568
  \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1569
  by (rule bit_eqI) (auto simp: bit_simps)
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
  1570
74497
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1571
lemma push_bit_Suc_minus_numeral [simp]:
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1572
  \<open>push_bit (Suc n) (- numeral k) = push_bit n (- numeral (Num.Bit0 k))\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1573
  using local.push_bit_Suc_numeral push_bit_minus by presburger
74497
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1574
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1575
lemma push_bit_minus_numeral [simp]:
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1576
  \<open>push_bit (numeral l) (- numeral k) = push_bit (pred_numeral l) (- numeral (Num.Bit0 k))\<close>
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1577
  by (simp only: numeral_eq_Suc push_bit_Suc_minus_numeral)
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  1578
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1579
lemma take_bit_Suc_minus_1_eq:
74498
27475e64a887 more complete simp rules
haftmann
parents: 74497
diff changeset
  1580
  \<open>take_bit (Suc n) (- 1) = 2 ^ Suc n - 1\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1581
  by (simp add: mask_eq_exp_minus_1)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1582
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1583
lemma take_bit_numeral_minus_1_eq:
74498
27475e64a887 more complete simp rules
haftmann
parents: 74497
diff changeset
  1584
  \<open>take_bit (numeral k) (- 1) = 2 ^ numeral k - 1\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1585
  by (simp add: mask_eq_exp_minus_1)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1586
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1587
lemma push_bit_mask_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1588
  \<open>push_bit m (mask n) = mask (n + m) AND NOT (mask m)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1589
  by (rule bit_eqI) (auto simp: bit_simps not_less possible_bit_less_imp)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1590
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1591
lemma slice_eq_mask:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1592
  \<open>push_bit n (take_bit m (drop_bit n a)) = a AND mask (m + n) AND NOT (mask n)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1593
  by (rule bit_eqI) (auto simp: bit_simps)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1594
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1595
lemma push_bit_numeral_minus_1 [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1596
  \<open>push_bit (numeral n) (- 1) = - (2 ^ numeral n)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  1597
  by (simp add: push_bit_eq_mult)
74498
27475e64a887 more complete simp rules
haftmann
parents: 74497
diff changeset
  1598
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1599
lemma unset_bit_eq:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1600
  \<open>unset_bit n a = a - of_bool (bit a n) * 2 ^ n\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1601
proof -
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1602
  have \<open>NOT a AND of_bool (bit a n) * 2 ^ n = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1603
    by (auto simp: bit_eq_iff bit_simps)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1604
  moreover have \<open>unset_bit n a = a AND NOT (of_bool (bit a n) * 2 ^ n)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1605
    by (auto simp: bit_eq_iff bit_simps)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1606
  ultimately show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  1607
    by (simp add: disjunctive_diff_eq_and_not)
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1608
qed
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1609
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1610
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1611
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1612
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1613
subsection \<open>Common algebraic structure\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1614
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1615
class linordered_euclidean_semiring_bit_operations =
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1616
  linordered_euclidean_semiring + semiring_bit_operations
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1617
begin
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1618
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1619
lemma possible_bit [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1620
  \<open>possible_bit TYPE('a) n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1621
  by (simp add: possible_bit_def)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1622
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1623
lemma take_bit_of_exp [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1624
  \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1625
  by (simp add: take_bit_eq_mod exp_mod_exp)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1626
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1627
lemma take_bit_of_2 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1628
  \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1629
  using take_bit_of_exp [of n 1] by simp
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1630
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1631
lemma push_bit_eq_0_iff [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1632
  \<open>push_bit n a = 0 \<longleftrightarrow> a = 0\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1633
  by (simp add: push_bit_eq_mult)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1634
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1635
lemma take_bit_add:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1636
  \<open>take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1637
  by (simp add: take_bit_eq_mod mod_simps)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1638
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1639
lemma take_bit_of_1_eq_0_iff [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1640
  \<open>take_bit n 1 = 0 \<longleftrightarrow> n = 0\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1641
  by (simp add: take_bit_eq_mod)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1642
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1643
lemma drop_bit_Suc_bit0 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1644
  \<open>drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1645
  by (simp add: drop_bit_Suc numeral_Bit0_div_2)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1646
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1647
lemma drop_bit_Suc_bit1 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1648
  \<open>drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\<close>
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1649
  by (simp add: drop_bit_Suc numeral_Bit0_div_2)
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1650
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1651
lemma drop_bit_numeral_bit0 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1652
  \<open>drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1653
  by (simp add: drop_bit_rec numeral_Bit0_div_2)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1654
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1655
lemma drop_bit_numeral_bit1 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1656
  \<open>drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1657
  by (simp add: drop_bit_rec numeral_Bit0_div_2)
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1658
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1659
lemma take_bit_Suc_1 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1660
  \<open>take_bit (Suc n) 1 = 1\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1661
  by (simp add: take_bit_Suc)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1662
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1663
lemma take_bit_Suc_bit0:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1664
  \<open>take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1665
  by (simp add: take_bit_Suc numeral_Bit0_div_2)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1666
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1667
lemma take_bit_Suc_bit1:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1668
  \<open>take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\<close>
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1669
  by (simp add: take_bit_Suc numeral_Bit0_div_2 mod_2_eq_odd)
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1670
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1671
lemma take_bit_numeral_1 [simp]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1672
  \<open>take_bit (numeral l) 1 = 1\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1673
  by (simp add: take_bit_rec [of \<open>numeral l\<close> 1])
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1674
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1675
lemma take_bit_numeral_bit0:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1676
  \<open>take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1677
  by (simp add: take_bit_rec numeral_Bit0_div_2)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1678
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1679
lemma take_bit_numeral_bit1:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1680
  \<open>take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\<close>
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  1681
  by (simp add: take_bit_rec numeral_Bit0_div_2 mod_2_eq_odd)
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1682
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1683
lemma bit_of_nat_iff_bit [bit_simps]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1684
  \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1685
proof -
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1686
  have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1687
    by simp
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1688
  also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1689
    by (simp add: of_nat_div)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1690
  finally show ?thesis
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1691
    by (simp add: bit_iff_odd semiring_bits_class.bit_iff_odd)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1692
qed
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1693
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1694
lemma drop_bit_mask_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1695
  \<open>drop_bit m (mask n) = mask (n - m)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1696
  by (rule bit_eqI) (auto simp: bit_simps possible_bit_def)
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1697
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1698
lemma bit_push_bit_iff':
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1699
  \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> bit a (n - m)\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1700
  by (simp add: bit_simps)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1701
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1702
lemma mask_half:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1703
  \<open>mask n div 2 = mask (n - 1)\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1704
  by (cases n) (simp_all add: mask_Suc_double one_or_eq)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1705
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1706
lemma take_bit_Suc_from_most:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1707
  \<open>take_bit (Suc n) a = 2 ^ n * of_bool (bit a n) + take_bit n a\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1708
  using mod_mult2_eq' [of a \<open>2 ^ n\<close> 2]
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1709
  by (simp only: take_bit_eq_mod power_Suc2)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1710
    (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1711
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1712
lemma take_bit_nonnegative [simp]:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1713
  \<open>0 \<le> take_bit n a\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1714
  using horner_sum_nonnegative by (simp flip: horner_sum_bit_eq_take_bit)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1715
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1716
lemma not_take_bit_negative [simp]:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1717
  \<open>\<not> take_bit n a < 0\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1718
  by (simp add: not_less)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1719
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1720
lemma bit_imp_take_bit_positive:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1721
  \<open>0 < take_bit m a\<close> if \<open>n < m\<close> and \<open>bit a n\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1722
proof (rule ccontr)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1723
  assume \<open>\<not> 0 < take_bit m a\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1724
  then have \<open>take_bit m a = 0\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1725
    by (auto simp: not_less intro: order_antisym)
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1726
  then have \<open>bit (take_bit m a) n = bit 0 n\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1727
    by simp
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1728
  with that show False
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1729
    by (simp add: bit_take_bit_iff)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1730
qed
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1731
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1732
lemma take_bit_mult:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1733
  \<open>take_bit n (take_bit n a * take_bit n b) = take_bit n (a * b)\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1734
  by (simp add: take_bit_eq_mod mod_mult_eq)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1735
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1736
lemma drop_bit_push_bit:
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1737
  \<open>drop_bit m (push_bit n a) = drop_bit (m - n) (push_bit (n - m) a)\<close>
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1738
  by (cases \<open>m \<le> n\<close>)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1739
    (auto simp: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] mult.assoc
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1740
    mult.commute [of a] drop_bit_eq_div push_bit_eq_mult not_le power_add Orderings.not_le dest!: le_Suc_ex less_imp_Suc_add)
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1741
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1742
end
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1743
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1744
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  1745
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
  1746
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1747
locale fold2_bit_int =
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1748
  fixes f :: \<open>bool \<Rightarrow> bool \<Rightarrow> bool\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1749
begin
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1750
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1751
context
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1752
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1753
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1754
function F :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1755
  where \<open>F k l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1756
    then - of_bool (f (odd k) (odd l))
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1757
    else of_bool (f (odd k) (odd l)) + 2 * (F (k div 2) (l div 2)))\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1758
  by auto
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1759
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1760
private termination proof (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>)
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1761
  have less_eq: \<open>\<bar>k div 2\<bar> \<le> \<bar>k\<bar>\<close> for k :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1762
    by (cases k) (simp_all add: divide_int_def nat_add_distrib)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  1763
  then have less: \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close> if \<open>k \<notin> {0, - 1}\<close> for k :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1764
    using that by (auto simp: less_le [of k])
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1765
  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
  1766
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1767
  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
  1768
    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
  1769
  proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1770
    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
  1771
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1772
    then have \<open>0 < \<bar>k\<bar> + \<bar>l\<bar>\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1773
      by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1774
    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
  1775
    proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1776
      assume \<open>k \<notin> {0, - 1}\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1777
      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
  1778
        by (rule less)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1779
      with less_eq [of l] show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1780
        by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1781
    next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1782
      assume \<open>l \<notin> {0, - 1}\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1783
      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
  1784
        by (rule less)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1785
      with less_eq [of k] show ?thesis
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1786
        by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1787
    qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1788
    ultimately show ?thesis
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1789
      by (simp only: in_measure split_def fst_conv snd_conv nat_mono_iff)
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1790
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  1791
qed
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1792
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1793
declare F.simps [simp del]
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1794
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1795
lemma rec:
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1796
  \<open>F k l = of_bool (f (odd k) (odd l)) + 2 * (F (k div 2) (l div 2))\<close>
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1797
    for k l :: int
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1798
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
  1799
  case True
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1800
  then show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1801
    by (auto simp: F.simps [of 0] F.simps [of \<open>- 1\<close>])
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1802
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1803
  case False
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1804
  then show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1805
    by (auto simp: ac_simps F.simps [of k l])
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1806
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1807
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1808
lemma bit_iff:
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1809
  \<open>bit (F k l) n \<longleftrightarrow> f (bit k n) (bit l n)\<close> for k l :: int
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1810
proof (induction n arbitrary: k l)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1811
  case 0
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1812
  then show ?case
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1813
    by (simp add: rec [of k l] bit_0)
71804
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1814
next
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1815
  case (Suc n)
6fd70ed18199 simplified construction of binary bit operations
haftmann
parents: 71802
diff changeset
  1816
  then show ?case
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1817
    by (simp add: rec [of k l] bit_Suc)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1818
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1819
79030
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1820
end
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1821
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1822
end
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1823
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1824
instantiation int :: ring_bit_operations
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1825
begin
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1826
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1827
definition not_int :: \<open>int \<Rightarrow> int\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1828
  where \<open>not_int k = - k - 1\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1829
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1830
global_interpretation and_int: fold2_bit_int \<open>(\<and>)\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1831
  defines and_int = and_int.F .
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1832
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1833
global_interpretation or_int: fold2_bit_int \<open>(\<or>)\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1834
  defines or_int = or_int.F .
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1835
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1836
global_interpretation xor_int: fold2_bit_int \<open>(\<noteq>)\<close>
bdea2b95817b more direct characterization of binary bit operations
haftmann
parents: 79018
diff changeset
  1837
  defines xor_int = xor_int.F .
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  1838
72082
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1839
definition mask_int :: \<open>nat \<Rightarrow> int\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1840
  where \<open>mask n = (2 :: int) ^ n - 1\<close>
41393ecb57ac uniform mask operation
haftmann
parents: 72079
diff changeset
  1841
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1842
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
  1843
  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
  1844
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1845
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
  1846
  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
  1847
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1848
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
  1849
  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
  1850
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1851
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
  1852
  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
  1853
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1854
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
  1855
  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
  1856
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1857
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
  1858
  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
  1859
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1860
lemma not_int_div_2:
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1861
  \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1862
  by (simp add: not_int_def)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1863
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  1864
lemma bit_not_int_iff:
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1865
  \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close> for k :: int
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1866
proof (rule sym, induction n arbitrary: k)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1867
  case 0
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1868
  then show ?case
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1869
    by (simp add: bit_0 not_int_def)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1870
next
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1871
  case (Suc n)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1872
  then show ?case
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1873
    by (simp add: bit_Suc not_int_div_2)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1874
qed
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  1875
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1876
instance proof
73682
78044b2f001c explicit type class operations for type-specific implementations
haftmann
parents: 73535
diff changeset
  1877
  fix k l :: int and m n :: nat
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1878
  show \<open>unset_bit n k = (k OR push_bit n 1) XOR push_bit n 1\<close>
79031
4596a14d9a95 slightly more elementary characterization of unset_bit
haftmann
parents: 79030
diff changeset
  1879
    by (rule bit_eqI)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1880
      (auto simp: unset_bit_int_def
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  1881
        and_int.bit_iff or_int.bit_iff xor_int.bit_iff bit_not_int_iff push_bit_int_def bit_simps)
79031
4596a14d9a95 slightly more elementary characterization of unset_bit
haftmann
parents: 79030
diff changeset
  1882
qed (fact and_int.rec or_int.rec xor_int.rec mask_int_def set_bit_int_def flip_bit_int_def
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  1883
  push_bit_int_def drop_bit_int_def take_bit_int_def not_int_def)+
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1884
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1885
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
  1886
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1887
instance int :: linordered_euclidean_semiring_bit_operations ..
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1888
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1889
context ring_bit_operations
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1890
begin
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1891
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1892
lemma even_of_int_iff:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1893
  \<open>even (of_int k) \<longleftrightarrow> even k\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1894
  by (induction k rule: int_bit_induct) simp_all
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1895
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1896
lemma bit_of_int_iff [bit_simps]:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1897
  \<open>bit (of_int k) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit k n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1898
proof (cases \<open>possible_bit TYPE('a) n\<close>)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1899
  case False
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1900
  then show ?thesis
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1901
    by (simp add: impossible_bit)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1902
next
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1903
  case True
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1904
  then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1905
  proof (induction k arbitrary: n rule: int_bit_induct)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1906
    case zero
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1907
    then show ?case
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1908
      by simp
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1909
  next
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1910
    case minus
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1911
    then show ?case
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1912
      by simp
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1913
  next
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1914
    case (even k)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1915
    then show ?case
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1916
      using bit_double_iff [of \<open>of_int k\<close> n] Bit_Operations.bit_double_iff [of k n]
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1917
      by (cases n) (auto simp: ac_simps possible_bit_def dest: mult_not_zero)
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1918
  next
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1919
    case (odd k)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1920
    then show ?case
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1921
      using bit_double_iff [of \<open>of_int k\<close> n]
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1922
      by (cases n)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  1923
        (auto simp: ac_simps bit_double_iff even_bit_succ_iff Bit_Operations.bit_0 Bit_Operations.bit_Suc
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1924
          possible_bit_def dest: mult_not_zero)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1925
  qed
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1926
  with True show ?thesis
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1927
    by simp
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1928
qed
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1929
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1930
lemma push_bit_of_int:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1931
  \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1932
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1933
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1934
lemma of_int_push_bit:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1935
  \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1936
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1937
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1938
lemma take_bit_of_int:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1939
  \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1940
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1941
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1942
lemma of_int_take_bit:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1943
  \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1944
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1945
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1946
lemma of_int_not_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1947
  \<open>of_int (NOT k) = NOT (of_int k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1948
  by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1949
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1950
lemma of_int_not_numeral:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1951
  \<open>of_int (NOT (numeral k)) = NOT (numeral k)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1952
  by (simp add: local.of_int_not_eq)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1953
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1954
lemma of_int_and_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1955
  \<open>of_int (k AND l) = of_int k AND of_int l\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1956
  by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1957
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1958
lemma of_int_or_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1959
  \<open>of_int (k OR l) = of_int k OR of_int l\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1960
  by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1961
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1962
lemma of_int_xor_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1963
  \<open>of_int (k XOR l) = of_int k XOR of_int l\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1964
  by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1965
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1966
lemma of_int_mask_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1967
  \<open>of_int (mask n) = mask n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1968
  by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1969
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1970
end
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  1971
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1972
lemma take_bit_int_less_exp [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1973
  \<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
  1974
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1975
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1976
lemma take_bit_int_eq_self_iff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1977
  \<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
  1978
  for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1979
proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1980
  assume ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1981
  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
  1982
  ultimately show ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1983
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1984
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1985
  assume ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1986
  then show ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1987
    by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1988
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1989
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1990
lemma take_bit_int_eq_self:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  1991
  \<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
  1992
  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
  1993
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1994
lemma mask_nonnegative_int [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1995
  \<open>mask n \<ge> (0::int)\<close>
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  1996
  by (simp add: mask_eq_exp_minus_1 add_le_imp_le_diff)
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1997
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1998
lemma not_mask_negative_int [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  1999
  \<open>\<not> mask n < (0::int)\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2000
  by (simp add: not_less)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  2001
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2002
lemma not_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2003
  \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2004
  by (simp add: not_int_def)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2005
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2006
lemma not_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2007
  \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2008
  by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2009
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2010
lemma and_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2011
  \<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
  2012
proof (induction k arbitrary: l rule: int_bit_induct)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2013
  case zero
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2014
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2015
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2016
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2017
  case minus
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2018
  then show ?case
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2019
    by simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2020
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2021
  case (even k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2022
  then show ?case
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2023
    using and_int.rec [of \<open>k * 2\<close> l]
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2024
    by (simp add: pos_imp_zdiv_nonneg_iff zero_le_mult_iff)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2025
next
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2026
  case (odd k)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2027
  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
  2028
    by simp
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  2029
  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
  2030
    by simp
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2031
  with and_int.rec [of \<open>1 + k * 2\<close> l]
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2032
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2033
    by (auto simp: zero_le_mult_iff not_le)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2034
qed
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2035
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2036
lemma and_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2037
  \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2038
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2039
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2040
lemma and_less_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2041
  \<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
  2042
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2043
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2044
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2045
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2046
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2047
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2048
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2049
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2050
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2051
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2052
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2053
  show ?case
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2054
    by (simp add: and_int.rec [of _ l])
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2055
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2056
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2057
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2058
  show ?case
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2059
    by (simp add: and_int.rec [of _ l])
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2060
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2061
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2062
lemma or_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2063
  \<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
  2064
  by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2065
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2066
lemma or_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2067
  \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2068
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2069
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2070
lemma or_greater_eq:
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2071
  \<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
  2072
using that proof (induction k arbitrary: l rule: int_bit_induct)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2073
  case zero
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2074
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2075
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2076
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2077
  case minus
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2078
  then show ?case
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2079
    by simp
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2080
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2081
  case (even k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2082
  from even.IH [of \<open>l div 2\<close>] even.hyps even.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2083
  show ?case
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2084
    by (simp add: or_int.rec [of _ l])
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2085
next
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2086
  case (odd k)
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2087
  from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2088
  show ?case
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2089
    by (simp add: or_int.rec [of _ l])
72009
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2090
qed
febdd4eead56 more on single-bit operations
haftmann
parents: 71991
diff changeset
  2091
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2092
lemma xor_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2093
  \<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
  2094
  by (simp only: bit.xor_def or_nonnegative_int_iff) auto
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2095
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2096
lemma xor_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2097
  \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2098
  by (subst Not_eq_iff [symmetric]) (auto simp: not_less)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2099
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2100
lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2101
  \<open>x OR y < 2 ^ n\<close> if \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2102
using that proof (induction x arbitrary: y n rule: int_bit_induct)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2103
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2104
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2105
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2106
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2107
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2108
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2109
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2110
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2111
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2112
  from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2113
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2114
    by (cases n) (auto simp: or_int.rec [of \<open>_ * 2\<close>] elim: oddE)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2115
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2116
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2117
  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
  2118
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2119
    by (cases n) (auto simp: or_int.rec [of \<open>1 + _ * 2\<close>], linarith)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2120
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2121
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2122
lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2123
  \<open>x XOR y < 2 ^ n\<close> if \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2124
using that proof (induction x arbitrary: y n rule: int_bit_induct)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2125
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2126
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2127
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2128
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2129
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2130
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2131
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2132
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2133
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2134
  from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2135
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2136
    by (cases n) (auto simp: xor_int.rec [of \<open>_ * 2\<close>] elim: oddE)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2137
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2138
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2139
  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
  2140
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2141
    by (cases n) (auto simp: xor_int.rec [of \<open>1 + _ * 2\<close>])
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2142
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2143
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2144
lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2145
  \<open>0 \<le> x AND y\<close> if \<open>0 \<le> x\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2146
  using that by simp
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2147
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2148
lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2149
  \<open>0 \<le> x OR y\<close> if \<open>0 \<le> x\<close> \<open>0 \<le> y\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2150
  using that by simp
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2151
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2152
lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2153
  \<open>0 \<le> x XOR y\<close> if \<open>0 \<le> x\<close> \<open>0 \<le> y\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2154
  using that by simp
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2155
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2156
lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2157
  \<open>x AND y \<le> x\<close> if \<open>0 \<le> x\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2158
using that proof (induction x arbitrary: y rule: int_bit_induct)
73535
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  2159
  case (odd k)
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  2160
  then have \<open>k AND y div 2 \<le> k\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 72830
diff changeset
  2161
    by simp
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2162
  then show ?case
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2163
    by (simp add: and_int.rec [of \<open>1 + _ * 2\<close>])
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2164
qed (simp_all add: and_int.rec [of \<open>_ * 2\<close>])
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2165
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2166
lemma AND_upper1' [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2167
  \<open>y AND x \<le> z\<close> if \<open>0 \<le> y\<close> \<open>y \<le> z\<close> for x y z :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2168
  using _ \<open>y \<le> z\<close> by (rule order_trans) (use \<open>0 \<le> y\<close> in simp)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2169
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2170
lemma AND_upper1'' [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2171
  \<open>y AND x < z\<close> if \<open>0 \<le> y\<close> \<open>y < z\<close> for x y z :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2172
  using _ \<open>y < z\<close> by (rule order_le_less_trans) (use \<open>0 \<le> y\<close> in simp)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2173
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2174
lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2175
  \<open>x AND y \<le> y\<close> if \<open>0 \<le> y\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2176
  using that AND_upper1 [of y x] by (simp add: ac_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2177
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2178
lemma AND_upper2' [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2179
  \<open>x AND y \<le> z\<close> if \<open>0 \<le> y\<close> \<open>y \<le> z\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2180
  using that AND_upper1' [of y z x] by (simp add: ac_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2181
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2182
lemma AND_upper2'' [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2183
  \<open>x AND y < z\<close> if \<open>0 \<le> y\<close> \<open>y < z\<close> for x y :: int
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2184
  using that AND_upper1'' [of y z x] by (simp add: ac_simps)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2185
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2186
lemma plus_and_or:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2187
  \<open>(x AND y) + (x OR y) = x + y\<close> for x y :: int
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2188
proof (induction x arbitrary: y rule: int_bit_induct)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2189
  case zero
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2190
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2191
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2192
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2193
  case minus
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2194
  then show ?case
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2195
    by simp
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2196
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2197
  case (even x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2198
  from even.IH [of \<open>y div 2\<close>]
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2199
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2200
    by (auto simp: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2201
next
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2202
  case (odd x)
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2203
  from odd.IH [of \<open>y div 2\<close>]
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2204
  show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2205
    by (auto simp: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE)
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2206
qed
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  2207
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2208
lemma push_bit_minus_one:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2209
  \<open>push_bit n (- 1 :: int) = - (2 ^ n)\<close>
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2210
  by (simp add: push_bit_eq_mult)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2211
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2212
lemma minus_1_div_exp_eq_int:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2213
  \<open>- 1 div (2 :: int) ^ n = - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2214
  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
  2215
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2216
lemma drop_bit_minus_one [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2217
  \<open>drop_bit n (- 1 :: int) = - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2218
  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
  2219
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2220
lemma take_bit_minus:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2221
  \<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
  2222
    for k :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2223
  by (simp add: take_bit_eq_mod mod_minus_eq)
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
lemma take_bit_diff:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2226
  \<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
  2227
    for k l :: int
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2228
  by (simp add: take_bit_eq_mod mod_diff_eq)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2229
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2230
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
  2231
  \<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
  2232
  by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2233
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2234
lemma take_bit_minus_small_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2235
  \<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
  2236
proof -
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2237
  define m where \<open>m = nat k\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2238
  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
  2239
    by simp_all
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2240
  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
  2241
    using \<open>0 < m\<close> by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2242
  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
  2243
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2244
  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
  2245
    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
  2246
  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
  2247
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2248
  then show ?thesis
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2249
    by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2250
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2251
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2252
lemma push_bit_nonnegative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2253
  \<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
  2254
  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
  2255
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2256
lemma push_bit_negative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2257
  \<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
  2258
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2259
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2260
lemma drop_bit_nonnegative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2261
  \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2262
  by (induction n) (auto simp: drop_bit_Suc drop_bit_half)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2263
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2264
lemma drop_bit_negative_int_iff [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2265
  \<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
  2266
  by (subst Not_eq_iff [symmetric]) (simp add: not_less)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2267
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2268
lemma set_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2269
  \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2270
  by (simp add: set_bit_eq_or)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2271
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2272
lemma set_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2273
  \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2274
  by (simp add: set_bit_eq_or)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2275
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2276
lemma unset_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2277
  \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2278
  by (simp add: unset_bit_eq_and_not)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2279
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2280
lemma unset_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2281
  \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2282
  by (simp add: unset_bit_eq_and_not)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2283
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2284
lemma flip_bit_nonnegative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2285
  \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2286
  by (simp add: flip_bit_eq_xor)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2287
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2288
lemma flip_bit_negative_int_iff [simp]:
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2289
  \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2290
  by (simp add: flip_bit_eq_xor)
71802
ab3cecb836b5 more rules
haftmann
parents: 71800
diff changeset
  2291
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2292
lemma set_bit_greater_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2293
  \<open>set_bit n k \<ge> k\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2294
  by (simp add: set_bit_eq_or or_greater_eq)
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2295
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2296
lemma unset_bit_less_eq:
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2297
  \<open>unset_bit n k \<le> k\<close> for k :: int
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2298
  by (simp add: unset_bit_eq_and_not and_less_eq)
71986
76193dd4aec8 factored out ancient numeral representation
haftmann
parents: 71965
diff changeset
  2299
75651
f4116b7a6679 Move code lemmas for symbolic computation of bit operations on int to distribution.
haftmann
parents: 75138
diff changeset
  2300
lemma and_int_unfold:
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2301
  \<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
  2302
    else (k mod 2) * (l mod 2) + 2 * ((k div 2) AND (l div 2)))\<close> for k l :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2303
  by (auto simp: and_int.rec [of k l] zmult_eq_1_iff elim: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2304
75651
f4116b7a6679 Move code lemmas for symbolic computation of bit operations on int to distribution.
haftmann
parents: 75138
diff changeset
  2305
lemma or_int_unfold:
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2306
  \<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
  2307
    else max (k mod 2) (l mod 2) + 2 * ((k div 2) OR (l div 2)))\<close> for k l :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2308
  by (auto simp: or_int.rec [of k l] elim: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2309
75651
f4116b7a6679 Move code lemmas for symbolic computation of bit operations on int to distribution.
haftmann
parents: 75138
diff changeset
  2310
lemma xor_int_unfold:
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2311
  \<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
  2312
    else \<bar>k mod 2 - l mod 2\<bar> + 2 * ((k div 2) XOR (l div 2)))\<close> for k l :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2313
  by (auto simp: xor_int.rec [of k l] not_int_def elim!: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2314
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2315
lemma bit_minus_int_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2316
  \<open>bit (- k) n \<longleftrightarrow> bit (NOT (k - 1)) n\<close> for k :: int
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2317
  by (simp add: bit_simps)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2318
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2319
lemma take_bit_incr_eq:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2320
  \<open>take_bit n (k + 1) = 1 + take_bit n k\<close> if \<open>take_bit n k \<noteq> 2 ^ n - 1\<close> for k :: int
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2321
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2322
  from that have \<open>2 ^ n \<noteq> k mod 2 ^ n + 1\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2323
    by (simp add: take_bit_eq_mod)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2324
  moreover have \<open>k mod 2 ^ n < 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2325
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2326
  ultimately have *: \<open>k mod 2 ^ n + 1 < 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2327
    by linarith
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2328
  have \<open>(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2329
    by (simp add: mod_simps)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2330
  also have \<open>\<dots> = k mod 2 ^ n + 1\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2331
    using * by (simp add: zmod_trivial_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2332
  finally have \<open>(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\<close> .
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2333
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2334
    by (simp add: take_bit_eq_mod)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2335
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2336
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2337
lemma take_bit_decr_eq:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2338
  \<open>take_bit n (k - 1) = take_bit n k - 1\<close> if \<open>take_bit n k \<noteq> 0\<close> for k :: int
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2339
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2340
  from that have \<open>k mod 2 ^ n \<noteq> 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2341
    by (simp add: take_bit_eq_mod)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2342
  moreover have \<open>k mod 2 ^ n \<ge> 0\<close> \<open>k mod 2 ^ n < 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2343
    by simp_all
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2344
  ultimately have *: \<open>k mod 2 ^ n > 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2345
    by linarith
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2346
  have \<open>(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2347
    by (simp add: mod_simps)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2348
  also have \<open>\<dots> = k mod 2 ^ n - 1\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2349
    by (simp add: zmod_trivial_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2350
      (use \<open>k mod 2 ^ n < 2 ^ n\<close> * in linarith)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2351
  finally have \<open>(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\<close> .
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2352
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2353
    by (simp add: take_bit_eq_mod)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2354
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2355
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2356
lemma take_bit_int_greater_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2357
  \<open>k + 2 ^ n \<le> take_bit n k\<close> if \<open>k < 0\<close> for k :: int
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2358
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2359
  have \<open>k + 2 ^ n \<le> take_bit n (k + 2 ^ n)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2360
  proof (cases \<open>k > - (2 ^ n)\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2361
    case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2362
    then have \<open>k + 2 ^ n \<le> 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2363
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2364
    also note take_bit_nonnegative
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2365
    finally show ?thesis .
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2366
  next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2367
    case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2368
    with that have \<open>0 \<le> k + 2 ^ n\<close> and \<open>k + 2 ^ n < 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2369
      by simp_all
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2370
    then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2371
      by (simp only: take_bit_eq_mod mod_pos_pos_trivial)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2372
  qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2373
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2374
    by (simp add: take_bit_eq_mod)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2375
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2376
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2377
lemma take_bit_int_less_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2378
  \<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
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2379
  using that zmod_le_nonneg_dividend [of \<open>k - 2 ^ n\<close> \<open>2 ^ n\<close>]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2380
  by (simp add: take_bit_eq_mod)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2381
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2382
lemma take_bit_int_less_eq_self_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2383
  \<open>take_bit n k \<le> k \<longleftrightarrow> 0 \<le> k\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) for k :: int
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2384
proof
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2385
  assume ?P
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2386
  show ?Q
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2387
  proof (rule ccontr)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2388
    assume \<open>\<not> 0 \<le> k\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2389
    then have \<open>k < 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2390
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2391
    with \<open>?P\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2392
    have \<open>take_bit n k < 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2393
      by (rule le_less_trans)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2394
    then show False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2395
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2396
  qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2397
next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2398
  assume ?Q
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2399
  then show ?P
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2400
    by (simp add: take_bit_eq_mod zmod_le_nonneg_dividend)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2401
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2402
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2403
lemma take_bit_int_less_self_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2404
  \<open>take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close> for k :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2405
  by (auto simp: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2406
    intro: order_trans [of 0 \<open>2 ^ n\<close> k])
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2407
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2408
lemma take_bit_int_greater_self_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2409
  \<open>k < take_bit n k \<longleftrightarrow> k < 0\<close> for k :: int
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2410
  using take_bit_int_less_eq_self_iff [of n k] by auto
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2411
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2412
lemma take_bit_int_greater_eq_self_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2413
  \<open>k \<le> take_bit n k \<longleftrightarrow> k < 2 ^ n\<close> for k :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2414
  by (auto simp: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2415
    dest: sym not_sym intro: less_trans [of k 0 \<open>2 ^ n\<close>])
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2416
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2417
lemma take_bit_tightened_less_eq_int:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2418
  \<open>take_bit m k \<le> take_bit n k\<close> if \<open>m \<le> n\<close> for k :: int
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2419
proof -
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2420
  have \<open>take_bit m (take_bit n k) \<le> take_bit n k\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2421
    by (simp only: take_bit_int_less_eq_self_iff take_bit_nonnegative)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2422
  with that show ?thesis
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2423
    by simp
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2424
qed
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2425
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2426
lemma not_exp_less_eq_0_int [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2427
  \<open>\<not> 2 ^ n \<le> (0::int)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2428
  by (simp add: power_le_zero_eq)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2429
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2430
lemma int_bit_bound:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2431
  fixes k :: int
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2432
  obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2433
    and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2434
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2435
  obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2436
  proof (cases \<open>k \<ge> 0\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2437
    case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2438
    moreover from power_gt_expt [of 2 \<open>nat k\<close>]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2439
    have \<open>nat k < 2 ^ nat k\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2440
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2441
    then have \<open>int (nat k) < int (2 ^ nat k)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2442
      by (simp only: of_nat_less_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2443
    ultimately have *: \<open>k div 2 ^ nat k = 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2444
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2445
    show thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2446
    proof (rule that [of \<open>nat k\<close>])
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2447
      fix m
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2448
      assume \<open>nat k \<le> m\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2449
      then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2450
        by (auto simp: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2451
    qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2452
  next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2453
    case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2454
    moreover from power_gt_expt [of 2 \<open>nat (- k)\<close>]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2455
    have \<open>nat (- k) < 2 ^ nat (- k)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2456
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2457
    then have \<open>int (nat (- k)) < int (2 ^ nat (- k))\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2458
      by (simp only: of_nat_less_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2459
    ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2460
      by (subst div_pos_neg_trivial) simp_all
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2461
    then have *: \<open>k div 2 ^ nat (- k) = - 1\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2462
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2463
    show thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2464
    proof (rule that [of \<open>nat (- k)\<close>])
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2465
      fix m
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2466
      assume \<open>nat (- k) \<le> m\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2467
      then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2468
        by (auto simp: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2469
    qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2470
  qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2471
  show thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2472
  proof (cases \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2473
    case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2474
    then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2475
      by blast
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2476
    with True that [of 0] show thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2477
      by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2478
  next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2479
    case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2480
    then obtain r where **: \<open>bit k r \<noteq> bit k q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2481
      by blast
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2482
    have \<open>r < q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2483
      by (rule ccontr) (use * [of r] ** in simp)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2484
    define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2485
    moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2486
      using ** N_def \<open>r < q\<close> by auto
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2487
    moreover define n where \<open>n = Suc (Max N)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2488
    ultimately have \<dagger>: \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2489
      by (smt (verit) "*" Max_ge Suc_n_not_le_n linorder_not_less mem_Collect_eq not_less_eq_eq)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2490
    have \<open>bit k (Max N) \<noteq> bit k n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2491
      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)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2492
    with \<dagger> n_def that [of n] show thesis
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2493
      by fastforce
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2494
  qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2495
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2496
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2497
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2498
subsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2499
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2500
instantiation nat :: semiring_bit_operations
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2501
begin
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2502
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2503
definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2504
  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
  2505
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2506
definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2507
  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
  2508
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2509
definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2510
  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
  2511
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2512
definition mask_nat :: \<open>nat \<Rightarrow> nat\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2513
  where \<open>mask n = (2 :: nat) ^ n - 1\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2514
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2515
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
  2516
  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
  2517
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2518
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
  2519
  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
  2520
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2521
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
  2522
  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
  2523
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2524
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
  2525
  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
  2526
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2527
definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  2528
  where \<open>unset_bit m n = (n OR push_bit m 1) XOR push_bit m 1\<close> for m n :: nat
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2529
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2530
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
  2531
  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
  2532
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2533
instance proof
79031
4596a14d9a95 slightly more elementary characterization of unset_bit
haftmann
parents: 79030
diff changeset
  2534
  fix m n :: nat
79008
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  2535
  show \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * (m div 2 AND n div 2)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  2536
    by (simp add: and_nat_def and_rec [of \<open>int m\<close> \<open>int n\<close>] nat_add_distrib of_nat_div)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  2537
  show \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * (m div 2 OR n div 2)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  2538
    by (simp add: or_nat_def or_rec [of \<open>int m\<close> \<open>int n\<close>] nat_add_distrib of_nat_div)
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  2539
  show \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * (m div 2 XOR n div 2)\<close>
74a4776f7a22 operations AND, OR, XOR are specified by characteristic recursive equation
haftmann
parents: 78955
diff changeset
  2540
    by (simp add: xor_nat_def xor_rec [of \<open>int m\<close> \<open>int n\<close>] nat_add_distrib of_nat_div)
79489
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  2541
qed (simp_all add: mask_nat_def set_bit_nat_def unset_bit_nat_def flip_bit_nat_def
1e19abf373ac streamlined type class specification
haftmann
parents: 79488
diff changeset
  2542
  push_bit_nat_def drop_bit_nat_def take_bit_nat_def)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2543
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2544
end
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2545
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2546
instance nat :: linordered_euclidean_semiring_bit_operations ..
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2547
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2548
context semiring_bit_operations
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2549
begin
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2550
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2551
lemma push_bit_of_nat:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2552
  \<open>push_bit n (of_nat m) = of_nat (push_bit n m)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2553
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2554
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2555
lemma of_nat_push_bit:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2556
  \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2557
  by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2558
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2559
lemma take_bit_of_nat:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2560
  \<open>take_bit n (of_nat m) = of_nat (take_bit n m)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2561
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2562
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2563
lemma of_nat_take_bit:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2564
  \<open>of_nat (take_bit n m) = take_bit n (of_nat m)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2565
  by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2566
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2567
lemma of_nat_and_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2568
  \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2569
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2570
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2571
lemma of_nat_or_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2572
  \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2573
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2574
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2575
lemma of_nat_xor_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2576
  \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2577
  by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2578
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2579
lemma of_nat_mask_eq:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2580
  \<open>of_nat (mask n) = mask n\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2581
  by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2582
81722
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2583
lemma of_nat_set_bit_eq:
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2584
  \<open>of_nat (set_bit n m) = set_bit n (of_nat m)\<close>
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2585
  by (simp add: set_bit_eq_or Bit_Operations.set_bit_eq_or of_nat_or_eq Bit_Operations.push_bit_eq_mult)
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2586
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2587
lemma of_nat_unset_bit_eq:
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2588
  \<open>of_nat (unset_bit n m) = unset_bit n (of_nat m)\<close>
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2589
  by (simp add: unset_bit_eq_or_xor Bit_Operations.unset_bit_eq_or_xor of_nat_or_eq of_nat_xor_eq Bit_Operations.push_bit_eq_mult)
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2590
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2591
lemma of_nat_flip_bit_eq:
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2592
  \<open>of_nat (flip_bit n m) = flip_bit n (of_nat m)\<close>
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2593
  by (simp add: flip_bit_eq_xor Bit_Operations.flip_bit_eq_xor of_nat_xor_eq Bit_Operations.push_bit_eq_mult)
658f1b5168f2 delegate computation to integer thoroughly
haftmann
parents: 81641
diff changeset
  2594
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2595
end
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2596
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2597
context linordered_euclidean_semiring_bit_operations
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2598
begin
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2599
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2600
lemma drop_bit_of_nat:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2601
  "drop_bit n (of_nat m) = of_nat (drop_bit n m)"
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2602
  by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div [of m "2 ^ n"])
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2603
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2604
lemma of_nat_drop_bit:
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2605
  \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close>
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2606
  by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2607
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2608
end
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  2609
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2610
lemma take_bit_nat_less_exp [simp]:
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2611
  \<open>take_bit n m < 2 ^ n\<close> for n m :: nat
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2612
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2613
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2614
lemma take_bit_nat_eq_self_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2615
  \<open>take_bit n m = m \<longleftrightarrow> m < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) for n m :: nat
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2616
proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2617
  assume ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2618
  moreover note take_bit_nat_less_exp [of n m]
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2619
  ultimately show ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2620
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2621
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2622
  assume ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2623
  then show ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2624
    by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2625
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2626
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2627
lemma take_bit_nat_eq_self:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2628
  \<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
  2629
  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
  2630
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2631
lemma take_bit_nat_less_eq_self [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2632
  \<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
  2633
  by (simp add: take_bit_eq_mod)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2634
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2635
lemma take_bit_nat_less_self_iff:
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2636
  \<open>take_bit n m < m \<longleftrightarrow> 2 ^ n \<le> m\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) for m n :: nat
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2637
proof
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2638
  assume ?P
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2639
  then have \<open>take_bit n m \<noteq> m\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2640
    by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2641
  then show \<open>?Q\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2642
    by (simp add: take_bit_nat_eq_self_iff)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2643
next
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2644
  have \<open>take_bit n m < 2 ^ n\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2645
    by (fact take_bit_nat_less_exp)
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2646
  also assume ?Q
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2647
  finally show ?P .
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2648
qed
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2649
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2650
lemma Suc_0_and_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2651
  \<open>Suc 0 AND n = n mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2652
  using one_and_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2653
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2654
lemma and_Suc_0_eq [simp]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2655
  \<open>n AND Suc 0 = n mod 2\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2656
  using and_one_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2657
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2658
lemma Suc_0_or_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2659
  \<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
  2660
  using one_or_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2661
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2662
lemma or_Suc_0_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2663
  \<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
  2664
  using or_one_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2665
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2666
lemma Suc_0_xor_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2667
  \<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
  2668
  using one_xor_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2669
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2670
lemma xor_Suc_0_eq:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2671
  \<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
  2672
  using xor_one_eq [of n] by simp
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2673
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2674
lemma and_nat_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2675
  \<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
  2676
    for m n :: nat
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2677
  by (auto simp: and_rec [of m n] elim: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2678
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2679
lemma or_nat_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2680
  \<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
  2681
    else max (m mod 2) (n mod 2) + 2 * ((m div 2) OR (n div 2)))\<close> for m n :: nat
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2682
  by (auto simp: or_rec [of m n] elim: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2683
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2684
lemma xor_nat_unfold [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2685
  \<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
  2686
    else (m mod 2 + n mod 2) mod 2 + 2 * ((m div 2) XOR (n div 2)))\<close> for m n :: nat
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2687
  by (auto simp: xor_rec [of m n] elim!: oddE)
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2688
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2689
lemma [code]:
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2690
  \<open>unset_bit 0 m = 2 * (m div 2)\<close>
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2691
  \<open>unset_bit (Suc n) m = m mod 2 + 2 * unset_bit n (m div 2)\<close> for m n :: nat
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  2692
  by (simp_all add: unset_bit_0 unset_bit_Suc)
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2693
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2694
lemma push_bit_of_Suc_0 [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2695
  \<open>push_bit n (Suc 0) = 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2696
  using push_bit_of_1 [where ?'a = nat] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2697
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2698
lemma take_bit_of_Suc_0 [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2699
  \<open>take_bit n (Suc 0) = of_bool (0 < n)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2700
  using take_bit_of_1 [where ?'a = nat] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2701
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2702
lemma drop_bit_of_Suc_0 [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2703
  \<open>drop_bit n (Suc 0) = of_bool (n = 0)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2704
  using drop_bit_of_1 [where ?'a = nat] by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2705
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2706
lemma Suc_mask_eq_exp:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2707
  \<open>Suc (mask n) = 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2708
  by (simp add: mask_eq_exp_minus_1)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2709
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2710
lemma less_eq_mask:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2711
  \<open>n \<le> mask n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2712
  by (simp add: mask_eq_exp_minus_1 le_diff_conv2)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2713
    (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2714
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2715
lemma less_mask:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2716
  \<open>n < mask n\<close> if \<open>Suc 0 < n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2717
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2718
  define m where \<open>m = n - 2\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2719
  with that have *: \<open>n = m + 2\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2720
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2721
  have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2722
    by (induction m) simp_all
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2723
  then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2724
    by (simp add: Suc_mask_eq_exp)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2725
  then have \<open>m + 2 < mask (m + 2)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2726
    by (simp add: less_le)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2727
  with * show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2728
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2729
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2730
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2731
lemma mask_nat_less_exp [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2732
  \<open>(mask n :: nat) < 2 ^ n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2733
  by (simp add: mask_eq_exp_minus_1)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2734
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2735
lemma mask_nat_positive_iff [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2736
  \<open>(0::nat) < mask n \<longleftrightarrow> 0 < n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2737
proof (cases \<open>n = 0\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2738
  case True
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2739
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2740
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2741
next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2742
  case False
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2743
  then have \<open>0 < n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2744
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2745
  then have \<open>(0::nat) < mask n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2746
    using less_eq_mask [of n] by (rule order_less_le_trans)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2747
  with \<open>0 < n\<close> show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2748
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2749
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2750
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2751
lemma take_bit_tightened_less_eq_nat:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2752
  \<open>take_bit m q \<le> take_bit n q\<close> if \<open>m \<le> n\<close> for q :: nat
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2753
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2754
  have \<open>take_bit m (take_bit n q) \<le> take_bit n q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2755
    by (rule take_bit_nat_less_eq_self)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2756
  with that show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2757
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2758
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2759
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2760
lemma push_bit_nat_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2761
  \<open>push_bit n (nat k) = nat (push_bit n k)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2762
  by (cases \<open>k \<ge> 0\<close>) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2763
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2764
lemma drop_bit_nat_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2765
  \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2766
proof (cases \<open>k \<ge> 0\<close>)
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2767
  case True
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2768
  then show ?thesis
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2769
    by (metis drop_bit_of_nat int_nat_eq nat_int) 
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  2770
qed (simp add: nat_eq_iff2)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2771
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2772
lemma take_bit_nat_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2773
  \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2774
  using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2775
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2776
lemma nat_take_bit_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2777
  \<open>nat (take_bit n k) = take_bit n (nat k)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2778
  if \<open>k \<ge> 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2779
  using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2780
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2781
lemma nat_mask_eq:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2782
  \<open>nat (mask n) = mask n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2783
  by (simp add: nat_eq_iff of_nat_mask_eq)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  2784
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  2785
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2786
subsection \<open>Symbolic computations on numeral expressions\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2787
75138
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2788
context semiring_bits
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2789
begin
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2790
81641
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2791
lemma bit_1_0 [simp]:
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2792
  \<open>bit 1 0\<close>
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2793
  by (simp add: bit_0)
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2794
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2795
lemma not_bit_1_Suc [simp]:
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2796
  \<open>\<not> bit 1 (Suc n)\<close>
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2797
  by (simp add: bit_Suc)
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2798
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2799
lemma not_bit_1_numeral [simp]:
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2800
  \<open>\<not> bit 1 (numeral m)\<close>
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2801
  by (simp add: numeral_eq_Suc)
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  2802
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2803
lemma not_bit_numeral_Bit0_0 [simp]:
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2804
  \<open>\<not> bit (numeral (Num.Bit0 m)) 0\<close>
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2805
  by (simp add: bit_0)
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2806
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2807
lemma bit_numeral_Bit1_0 [simp]:
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2808
  \<open>bit (numeral (Num.Bit1 m)) 0\<close>
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2809
  by (simp add: bit_0)
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2810
79590
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2811
lemma bit_numeral_Bit0_iff:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2812
  \<open>bit (numeral (num.Bit0 m)) n
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2813
    \<longleftrightarrow> possible_bit TYPE('a) n \<and> n > 0 \<and> bit (numeral m) (n - 1)\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2814
  by (simp only: numeral_Bit0_eq_double [of m] bit_simps) simp
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2815
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2816
lemma bit_numeral_Bit1_Suc_iff:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2817
  \<open>bit (numeral (num.Bit1 m)) (Suc n)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2818
    \<longleftrightarrow> possible_bit TYPE('a) (Suc n) \<and> bit (numeral m) n\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2819
  using even_bit_succ_iff [of \<open>2 * numeral m\<close> \<open>Suc n\<close>]
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2820
  by (simp only: numeral_Bit1_eq_inc_double [of m] bit_simps) simp
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2821
75138
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2822
end
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2823
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2824
context ring_bit_operations
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2825
begin
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2826
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2827
lemma not_bit_minus_numeral_Bit0_0 [simp]:
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2828
  \<open>\<not> bit (- numeral (Num.Bit0 m)) 0\<close>
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2829
  by (simp add: bit_0)
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2830
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2831
lemma bit_minus_numeral_Bit1_0 [simp]:
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2832
  \<open>bit (- numeral (Num.Bit1 m)) 0\<close>
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2833
  by (simp add: bit_0)
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2834
79590
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2835
lemma bit_minus_numeral_Bit0_Suc_iff:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2836
  \<open>bit (- numeral (num.Bit0 m)) (Suc n)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2837
    \<longleftrightarrow> possible_bit TYPE('a) (Suc n) \<and> bit (- numeral m) n\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2838
  by (simp only: numeral_Bit0_eq_double [of m] minus_mult_right bit_simps) auto
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2839
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2840
lemma bit_minus_numeral_Bit1_Suc_iff:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2841
  \<open>bit (- numeral (num.Bit1 m)) (Suc n)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2842
    \<longleftrightarrow> possible_bit TYPE('a) (Suc n) \<and> \<not> bit (numeral m) n\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2843
  by (simp only: numeral_Bit1_eq_inc_double [of m] minus_add_distrib minus_mult_right add_uminus_conv_diff
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2844
    bit_decr_iff bit_double_iff)
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2845
    auto
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2846
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2847
lemma bit_numeral_BitM_0 [simp]:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2848
  \<open>bit (numeral (Num.BitM m)) 0\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2849
  by (simp only: numeral_BitM bit_decr_iff not_bit_minus_numeral_Bit0_0) simp
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2850
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2851
lemma bit_numeral_BitM_Suc_iff:
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2852
  \<open>bit (numeral (Num.BitM m)) (Suc n) \<longleftrightarrow> possible_bit TYPE('a) (Suc n) \<and> \<not> bit (- numeral m) n\<close>
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2853
  by (simp_all only: numeral_BitM bit_decr_iff bit_minus_numeral_Bit0_Suc_iff) auto
b14c4cb37d99 more lemmas
haftmann
parents: 79588
diff changeset
  2854
75138
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2855
end
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2856
78955
74147aa81dbb more specific name for type class
haftmann
parents: 78937
diff changeset
  2857
context linordered_euclidean_semiring_bit_operations
75138
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2858
begin
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2859
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2860
lemma bit_numeral_iff:
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2861
  \<open>bit (numeral m) n \<longleftrightarrow> bit (numeral m :: nat) n\<close>
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2862
  using bit_of_nat_iff_bit [of \<open>numeral m\<close> n] by simp
cd77ffb01e15 simp rules for negative numerals
haftmann
parents: 75086
diff changeset
  2863
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2864
lemma bit_numeral_Bit0_Suc_iff [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2865
  \<open>bit (numeral (Num.Bit0 m)) (Suc n) \<longleftrightarrow> bit (numeral m) n\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2866
  by (simp add: bit_Suc numeral_Bit0_div_2)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2867
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2868
lemma bit_numeral_Bit1_Suc_iff [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2869
  \<open>bit (numeral (Num.Bit1 m)) (Suc n) \<longleftrightarrow> bit (numeral m) n\<close>
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  2870
  by (simp add: bit_Suc numeral_Bit0_div_2)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2871
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2872
lemma bit_numeral_rec:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2873
  \<open>bit (numeral (Num.Bit0 w)) n \<longleftrightarrow> (case n of 0 \<Rightarrow> False | Suc m \<Rightarrow> bit (numeral w) m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2874
  \<open>bit (numeral (Num.Bit1 w)) n \<longleftrightarrow> (case n of 0 \<Rightarrow> True | Suc m \<Rightarrow> bit (numeral w) m)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2875
  by (cases n; simp add: bit_0)+
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2876
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2877
lemma bit_numeral_simps [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2878
  \<open>bit (numeral (Num.Bit0 w)) (numeral n) \<longleftrightarrow> bit (numeral w) (pred_numeral n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2879
  \<open>bit (numeral (Num.Bit1 w)) (numeral n) \<longleftrightarrow> bit (numeral w) (pred_numeral n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2880
  by (simp_all add: bit_1_iff numeral_eq_Suc)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2881
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2882
lemma and_numerals [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2883
  \<open>1 AND numeral (Num.Bit0 y) = 0\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2884
  \<open>1 AND numeral (Num.Bit1 y) = 1\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2885
  \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = 2 * (numeral x AND numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2886
  \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = 2 * (numeral x AND numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2887
  \<open>numeral (Num.Bit0 x) AND 1 = 0\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2888
  \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = 2 * (numeral x AND numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2889
  \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + 2 * (numeral x AND numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2890
  \<open>numeral (Num.Bit1 x) AND 1 = 1\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2891
  by (simp_all add: bit_eq_iff) (simp_all add: bit_0 bit_simps bit_Suc bit_numeral_rec split: nat.splits)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2892
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2893
lemma or_numerals [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2894
  \<open>1 OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2895
  \<open>1 OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2896
  \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = 2 * (numeral x OR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2897
  \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + 2 * (numeral x OR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2898
  \<open>numeral (Num.Bit0 x) OR 1 = numeral (Num.Bit1 x)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2899
  \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + 2 * (numeral x OR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2900
  \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + 2 * (numeral x OR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2901
  \<open>numeral (Num.Bit1 x) OR 1 = numeral (Num.Bit1 x)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2902
  by (simp_all add: bit_eq_iff) (simp_all add: bit_0 bit_simps bit_Suc bit_numeral_rec split: nat.splits)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2903
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2904
lemma xor_numerals [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2905
  \<open>1 XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2906
  \<open>1 XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2907
  \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = 2 * (numeral x XOR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2908
  \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + 2 * (numeral x XOR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2909
  \<open>numeral (Num.Bit0 x) XOR 1 = numeral (Num.Bit1 x)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2910
  \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + 2 * (numeral x XOR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2911
  \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = 2 * (numeral x XOR numeral y)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2912
  \<open>numeral (Num.Bit1 x) XOR 1 = numeral (Num.Bit0 x)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  2913
  by (simp_all add: bit_eq_iff) (simp_all add: bit_0 bit_simps bit_Suc bit_numeral_rec split: nat.splits)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2914
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2915
end
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2916
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2917
lemma drop_bit_Suc_minus_bit0 [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2918
  \<open>drop_bit (Suc n) (- numeral (Num.Bit0 k)) = drop_bit n (- numeral k :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2919
  by (simp add: drop_bit_Suc numeral_Bit0_div_2)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2920
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2921
lemma drop_bit_Suc_minus_bit1 [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2922
  \<open>drop_bit (Suc n) (- numeral (Num.Bit1 k)) = drop_bit n (- numeral (Num.inc k) :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2923
  by (simp add: drop_bit_Suc numeral_Bit1_div_2 add_One)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2924
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2925
lemma drop_bit_numeral_minus_bit0 [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2926
  \<open>drop_bit (numeral l) (- numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (- numeral k :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2927
  by (simp add: numeral_eq_Suc numeral_Bit0_div_2)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2928
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2929
lemma drop_bit_numeral_minus_bit1 [simp]:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2930
  \<open>drop_bit (numeral l) (- numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (- numeral (Num.inc k) :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2931
  by (simp add: numeral_eq_Suc numeral_Bit1_div_2)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2932
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2933
lemma take_bit_Suc_minus_bit0:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2934
  \<open>take_bit (Suc n) (- numeral (Num.Bit0 k)) = take_bit n (- numeral k) * (2 :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2935
  by (simp add: take_bit_Suc numeral_Bit0_div_2)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2936
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2937
lemma take_bit_Suc_minus_bit1:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2938
  \<open>take_bit (Suc n) (- numeral (Num.Bit1 k)) = take_bit n (- numeral (Num.inc k)) * 2 + (1 :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2939
  by (simp add: take_bit_Suc numeral_Bit1_div_2 add_One)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2940
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2941
lemma take_bit_numeral_minus_bit0:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2942
  \<open>take_bit (numeral l) (- numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2943
  by (simp add: numeral_eq_Suc numeral_Bit0_div_2 take_bit_Suc_minus_bit0)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2944
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2945
lemma take_bit_numeral_minus_bit1:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2946
  \<open>take_bit (numeral l) (- numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (- numeral (Num.inc k)) * 2 + (1 :: int)\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2947
  by (simp add: numeral_eq_Suc numeral_Bit1_div_2 take_bit_Suc_minus_bit1)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 79008
diff changeset
  2948
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2949
lemma and_nat_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2950
  \<open>Suc 0 AND numeral (Num.Bit0 y) = 0\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2951
  \<open>Suc 0 AND numeral (Num.Bit1 y) = 1\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2952
  \<open>numeral (Num.Bit0 x) AND Suc 0 = 0\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2953
  \<open>numeral (Num.Bit1 x) AND Suc 0 = 1\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2954
  by (simp_all only: and_numerals flip: One_nat_def)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2955
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2956
lemma or_nat_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2957
  \<open>Suc 0 OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2958
  \<open>Suc 0 OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2959
  \<open>numeral (Num.Bit0 x) OR Suc 0 = numeral (Num.Bit1 x)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2960
  \<open>numeral (Num.Bit1 x) OR Suc 0 = numeral (Num.Bit1 x)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2961
  by (simp_all only: or_numerals flip: One_nat_def)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2962
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2963
lemma xor_nat_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2964
  \<open>Suc 0 XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2965
  \<open>Suc 0 XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2966
  \<open>numeral (Num.Bit0 x) XOR Suc 0 = numeral (Num.Bit1 x)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2967
  \<open>numeral (Num.Bit1 x) XOR Suc 0 = numeral (Num.Bit0 x)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2968
  by (simp_all only: xor_numerals flip: One_nat_def)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2969
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2970
context ring_bit_operations
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2971
begin
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2972
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2973
lemma minus_numeral_inc_eq:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2974
  \<open>- numeral (Num.inc n) = NOT (numeral n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2975
  by (simp add: not_eq_complement sub_inc_One_eq add_One)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2976
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2977
lemma sub_one_eq_not_neg:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2978
  \<open>Num.sub n num.One = NOT (- numeral n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2979
  by (simp add: not_eq_complement)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2980
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2981
lemma minus_numeral_eq_not_sub_one:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2982
  \<open>- numeral n = NOT (Num.sub n num.One)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2983
  by (simp add: not_eq_complement)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2984
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2985
lemma not_numeral_eq [simp]:
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2986
  \<open>NOT (numeral n) = - numeral (Num.inc n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2987
  by (simp add: minus_numeral_inc_eq)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2988
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2989
lemma not_minus_numeral_eq [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2990
  \<open>NOT (- numeral n) = Num.sub n num.One\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2991
  by (simp add: sub_one_eq_not_neg)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2992
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2993
lemma minus_not_numeral_eq [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  2994
  \<open>- (NOT (numeral n)) = numeral (Num.inc n)\<close>
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2995
  by simp
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2996
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2997
lemma not_numeral_BitM_eq:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  2998
  \<open>NOT (numeral (Num.BitM n)) =  - numeral (num.Bit0 n)\<close>
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  2999
  by (simp add: inc_BitM_eq)
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3000
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3001
lemma not_numeral_Bit0_eq:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3002
  \<open>NOT (numeral (Num.Bit0 n)) =  - numeral (num.Bit1 n)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3003
  by simp
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3004
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3005
end
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3006
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3007
lemma bit_minus_numeral_int [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3008
  \<open>bit (- numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (- numeral w :: int) (pred_numeral n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3009
  \<open>bit (- numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (numeral w :: int) (pred_numeral n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3010
  by (simp_all add: bit_minus_iff bit_not_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3011
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3012
lemma bit_minus_numeral_Bit0_Suc_iff [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3013
  \<open>bit (- numeral (num.Bit0 w) :: int) (Suc n) \<longleftrightarrow> bit (- numeral w :: int) n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3014
  by (simp add: bit_Suc)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3015
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3016
lemma bit_minus_numeral_Bit1_Suc_iff [simp]:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3017
  \<open>bit (- numeral (num.Bit1 w) :: int) (Suc n) \<longleftrightarrow> \<not> bit (numeral w :: int) n\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3018
  by (simp add: bit_Suc add_One flip: bit_not_int_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3019
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3020
lemma and_not_numerals:
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3021
  \<open>1 AND NOT 1 = (0 :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3022
  \<open>1 AND NOT (numeral (Num.Bit0 n)) = (1 :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3023
  \<open>1 AND NOT (numeral (Num.Bit1 n)) = (0 :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3024
  \<open>numeral (Num.Bit0 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3025
  \<open>numeral (Num.Bit0 m) AND NOT (numeral (Num.Bit0 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3026
  \<open>numeral (Num.Bit0 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3027
  \<open>numeral (Num.Bit1 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3028
  \<open>numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m AND NOT (numeral n))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3029
  \<open>numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close>
81641
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  3030
  by (simp_all add: bit_eq_iff)
5af6a5e4343b avoid default simp rule which would produce strange recursive unfolding in presence of bit_eq_iff
haftmann
parents: 81132
diff changeset
  3031
    (auto simp: bit_0 bit_simps bit_Suc bit_numeral_rec BitM_inc_eq sub_inc_One_eq split: nat.split)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3032
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3033
fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3034
where
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3035
  \<open>and_not_num num.One num.One = None\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3036
| \<open>and_not_num num.One (num.Bit0 n) = Some num.One\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3037
| \<open>and_not_num num.One (num.Bit1 n) = None\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3038
| \<open>and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3039
| \<open>and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3040
| \<open>and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3041
| \<open>and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3042
| \<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>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3043
| \<open>and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3044
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3045
lemma int_numeral_and_not_num:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3046
  \<open>numeral m AND NOT (numeral n) = (case and_not_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3047
  by (induction m n rule: and_not_num.induct) (simp_all del: not_numeral_eq not_one_eq add: and_not_numerals split: option.splits)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3048
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3049
lemma int_numeral_not_and_num:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3050
  \<open>NOT (numeral m) AND numeral n = (case and_not_num n m of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3051
  using int_numeral_and_not_num [of n m] by (simp add: ac_simps)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3052
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3053
lemma and_not_num_eq_None_iff:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3054
  \<open>and_not_num m n = None \<longleftrightarrow> numeral m AND NOT (numeral n) = (0 :: int)\<close>
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3055
  by (simp del: not_numeral_eq add: int_numeral_and_not_num split: option.split)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3056
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3057
lemma and_not_num_eq_Some_iff:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3058
  \<open>and_not_num m n = Some q \<longleftrightarrow> numeral m AND NOT (numeral n) = (numeral q :: int)\<close>
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3059
  by (simp del: not_numeral_eq add: int_numeral_and_not_num split: option.split)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3060
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3061
lemma and_minus_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3062
  \<open>1 AND - (numeral (num.Bit0 n)) = (0::int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3063
  \<open>1 AND - (numeral (num.Bit1 n)) = (1::int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3064
  \<open>numeral m AND - (numeral (num.Bit0 n)) = (case and_not_num m (Num.BitM n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3065
  \<open>numeral m AND - (numeral (num.Bit1 n)) = (case and_not_num m (Num.Bit0 n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3066
  \<open>- (numeral (num.Bit0 n)) AND 1 = (0::int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3067
  \<open>- (numeral (num.Bit1 n)) AND 1 = (1::int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3068
  \<open>- (numeral (num.Bit0 n)) AND numeral m = (case and_not_num m (Num.BitM n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3069
  \<open>- (numeral (num.Bit1 n)) AND numeral m = (case and_not_num m (Num.Bit0 n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3070
  by (simp_all del: not_numeral_eq add: ac_simps
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3071
    and_not_numerals one_and_eq not_numeral_BitM_eq not_numeral_Bit0_eq and_not_num_eq_None_iff and_not_num_eq_Some_iff split: option.split)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3072
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3073
lemma and_minus_minus_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3074
  \<open>- (numeral m :: int) AND - (numeral n :: int) = NOT ((numeral m - 1) OR (numeral n - 1))\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3075
  by (simp add: minus_numeral_eq_not_sub_one)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3076
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3077
lemma or_not_numerals:
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3078
  \<open>1 OR NOT 1 = NOT (0 :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3079
  \<open>1 OR NOT (numeral (Num.Bit0 n)) = NOT (numeral (Num.Bit0 n) :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3080
  \<open>1 OR NOT (numeral (Num.Bit1 n)) = NOT (numeral (Num.Bit0 n) :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3081
  \<open>numeral (Num.Bit0 m) OR NOT (1 :: int) = NOT (1 :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3082
  \<open>numeral (Num.Bit0 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3083
  \<open>numeral (Num.Bit0 m) OR NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m OR NOT (numeral n))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3084
  \<open>numeral (Num.Bit1 m) OR NOT (1 :: int) = NOT (0 :: int)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3085
  \<open>numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3086
  \<open>numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit1 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close>
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3087
  by (simp_all add: bit_eq_iff)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3088
    (auto simp: bit_0 bit_simps bit_Suc bit_numeral_rec sub_inc_One_eq split: nat.split)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3089
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3090
fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3091
where
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3092
  \<open>or_not_num_neg num.One num.One = num.One\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3093
| \<open>or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3094
| \<open>or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3095
| \<open>or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3096
| \<open>or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3097
| \<open>or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3098
| \<open>or_not_num_neg (num.Bit1 n) num.One = num.One\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3099
| \<open>or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3100
| \<open>or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3101
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3102
lemma int_numeral_or_not_num_neg:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3103
  \<open>numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\<close>
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3104
  by (induction m n rule: or_not_num_neg.induct) (simp_all del: not_numeral_eq not_one_eq add: or_not_numerals, simp_all)
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3105
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3106
lemma int_numeral_not_or_num_neg:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3107
  \<open>NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3108
  using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3109
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3110
lemma numeral_or_not_num_eq:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3111
  \<open>numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3112
  using int_numeral_or_not_num_neg [of m n] by simp
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3113
74495
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3114
lemma or_minus_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3115
  \<open>1 OR - (numeral (num.Bit0 n)) = - (numeral (or_not_num_neg num.One (Num.BitM n)) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3116
  \<open>1 OR - (numeral (num.Bit1 n)) = - (numeral (num.Bit1 n) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3117
  \<open>numeral m OR - (numeral (num.Bit0 n)) = - (numeral (or_not_num_neg m (Num.BitM n)) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3118
  \<open>numeral m OR - (numeral (num.Bit1 n)) = - (numeral (or_not_num_neg m (Num.Bit0 n)) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3119
  \<open>- (numeral (num.Bit0 n)) OR 1 = - (numeral (or_not_num_neg num.One (Num.BitM n)) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3120
  \<open>- (numeral (num.Bit1 n)) OR 1 = - (numeral (num.Bit1 n) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3121
  \<open>- (numeral (num.Bit0 n)) OR numeral m = - (numeral (or_not_num_neg m (Num.BitM n)) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3122
  \<open>- (numeral (num.Bit1 n)) OR numeral m = - (numeral (or_not_num_neg m (Num.Bit0 n)) :: int)\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3123
  by (simp_all only: or.commute [of _ 1] or.commute [of _ \<open>numeral m\<close>]
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3124
    minus_numeral_eq_not_sub_one or_not_numerals
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3125
    numeral_or_not_num_eq arith_simps minus_minus numeral_One)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3126
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3127
lemma or_minus_minus_numerals [simp]:
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3128
  \<open>- (numeral m :: int) OR - (numeral n :: int) = NOT ((numeral m - 1) AND (numeral n - 1))\<close>
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3129
  by (simp add: minus_numeral_eq_not_sub_one)
bc27c490aaac normalizing NOT (numeral _) (again)
haftmann
parents: 74391
diff changeset
  3130
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3131
lemma xor_minus_numerals [simp]:
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3132
  \<open>- numeral n XOR k = NOT (neg_numeral_class.sub n num.One XOR k)\<close>
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3133
  \<open>k XOR - numeral n = NOT (k XOR (neg_numeral_class.sub n num.One))\<close> for k :: int
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3134
  by (simp_all add: minus_numeral_eq_not_sub_one)
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3135
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3136
definition take_bit_num :: \<open>nat \<Rightarrow> num \<Rightarrow> num option\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3137
  where \<open>take_bit_num n m =
75651
f4116b7a6679 Move code lemmas for symbolic computation of bit operations on int to distribution.
haftmann
parents: 75138
diff changeset
  3138
    (if take_bit n (numeral m :: nat) = 0 then None else Some (num_of_nat (take_bit n (numeral m :: nat))))\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3139
74618
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3140
lemma take_bit_num_simps:
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3141
  \<open>take_bit_num 0 m = None\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3142
  \<open>take_bit_num (Suc n) Num.One =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3143
    Some Num.One\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3144
  \<open>take_bit_num (Suc n) (Num.Bit0 m) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3145
    (case take_bit_num n m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q))\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3146
  \<open>take_bit_num (Suc n) (Num.Bit1 m) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3147
    Some (case take_bit_num n m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q)\<close>
74618
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3148
  \<open>take_bit_num (numeral r) Num.One =
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3149
    Some Num.One\<close>
74618
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3150
  \<open>take_bit_num (numeral r) (Num.Bit0 m) =
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3151
    (case take_bit_num (pred_numeral r) m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q))\<close>
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3152
  \<open>take_bit_num (numeral r) (Num.Bit1 m) =
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3153
    Some (case take_bit_num (pred_numeral r) m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3154
  by (auto simp: take_bit_num_def ac_simps mult_2 num_of_nat_double
74618
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3155
    take_bit_Suc_bit0 take_bit_Suc_bit1 take_bit_numeral_bit0 take_bit_numeral_bit1)
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3156
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3157
lemma take_bit_num_code [code]:
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3158
  \<comment> \<open>Ocaml-style pattern matching is more robust wrt. different representations of \<^typ>\<open>nat\<close>\<close>
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3159
  \<open>take_bit_num n m = (case (n, m)
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3160
   of (0, _) \<Rightarrow> None
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3161
    | (Suc n, Num.One) \<Rightarrow> Some Num.One
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3162
    | (Suc n, Num.Bit0 m) \<Rightarrow> (case take_bit_num n m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q))
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3163
    | (Suc n, Num.Bit1 m) \<Rightarrow> Some (case take_bit_num n m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q))\<close>
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3164
  by (cases n; cases m) (simp_all add: take_bit_num_simps)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3165
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3166
context semiring_bit_operations
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3167
begin
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3168
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3169
lemma take_bit_num_eq_None_imp:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3170
  \<open>take_bit m (numeral n) = 0\<close> if \<open>take_bit_num m n = None\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3171
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3172
  from that have \<open>take_bit m (numeral n :: nat) = 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3173
    by (simp add: take_bit_num_def split: if_splits)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3174
  then have \<open>of_nat (take_bit m (numeral n)) = of_nat 0\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3175
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3176
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3177
    by (simp add: of_nat_take_bit)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3178
qed
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3179
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3180
lemma take_bit_num_eq_Some_imp:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3181
  \<open>take_bit m (numeral n) = numeral q\<close> if \<open>take_bit_num m n = Some q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3182
proof -
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3183
  from that have \<open>take_bit m (numeral n :: nat) = numeral q\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3184
    by (auto simp: take_bit_num_def Num.numeral_num_of_nat_unfold split: if_splits)
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3185
  then have \<open>of_nat (take_bit m (numeral n)) = of_nat (numeral q)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3186
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3187
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3188
    by (simp add: of_nat_take_bit)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3189
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3190
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3191
lemma take_bit_numeral_numeral:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3192
  \<open>take_bit (numeral m) (numeral n) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3193
    (case take_bit_num (numeral m) n of None \<Rightarrow> 0 | Some q \<Rightarrow> numeral q)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3194
  by (auto split: option.split dest: take_bit_num_eq_None_imp take_bit_num_eq_Some_imp)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3195
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3196
end
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3197
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3198
lemma take_bit_numeral_minus_numeral_int:
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3199
  \<open>take_bit (numeral m) (- numeral n :: int) =
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3200
    (case take_bit_num (numeral m) n of None \<Rightarrow> 0 | Some q \<Rightarrow> take_bit (numeral m) (2 ^ numeral m - numeral q))\<close> (is \<open>?lhs = ?rhs\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3201
proof (cases \<open>take_bit_num (numeral m) n\<close>)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3202
  case None
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3203
  then show ?thesis
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3204
    by (auto dest: take_bit_num_eq_None_imp [where ?'a = int] simp add: take_bit_eq_0_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3205
next
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3206
  case (Some q)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3207
  then have q: \<open>take_bit (numeral m) (numeral n :: int) = numeral q\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3208
    by (auto dest: take_bit_num_eq_Some_imp)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3209
  let ?T = \<open>take_bit (numeral m) :: int \<Rightarrow> int\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3210
  have *: \<open>?T (2 ^ numeral m) = ?T (?T 0)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3211
    by (simp add: take_bit_eq_0_iff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3212
  have \<open>?lhs = ?T (0 - numeral n)\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3213
    by simp
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3214
  also have \<open>\<dots> = ?T (?T (?T 0) - ?T (?T (numeral n)))\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3215
    by (simp only: take_bit_diff)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3216
  also have \<open>\<dots> = ?T (2 ^ numeral m - ?T (numeral n))\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3217
    by (simp only: take_bit_diff flip: *)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3218
  also have \<open>\<dots> = ?rhs\<close>
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3219
    by (simp add: q Some)
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3220
  finally show ?thesis .
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3221
qed
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3222
74618
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3223
declare take_bit_num_simps [simp]
43142ac556e6 moved generic implementation into HOL-Main
haftmann
parents: 74592
diff changeset
  3224
  take_bit_numeral_numeral [simp]
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3225
  take_bit_numeral_minus_numeral_int [simp]
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3226
74163
afe3c8ae1624 consolidation of rules for bit operations
haftmann
parents: 74123
diff changeset
  3227
79069
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3228
subsection \<open>Symbolic computations for code generation\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3229
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3230
lemma bit_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3231
  \<open>bit (0::int)               n      \<longleftrightarrow> False\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3232
  \<open>bit (Int.Neg num.One)      n      \<longleftrightarrow> True\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3233
  \<open>bit (Int.Pos num.One)      0      \<longleftrightarrow> True\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3234
  \<open>bit (Int.Pos (num.Bit0 m)) 0      \<longleftrightarrow> False\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3235
  \<open>bit (Int.Pos (num.Bit1 m)) 0      \<longleftrightarrow> True\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3236
  \<open>bit (Int.Neg (num.Bit0 m)) 0      \<longleftrightarrow> False\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3237
  \<open>bit (Int.Neg (num.Bit1 m)) 0      \<longleftrightarrow> True\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3238
  \<open>bit (Int.Pos num.One)      (Suc n) \<longleftrightarrow> False\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3239
  \<open>bit (Int.Pos (num.Bit0 m)) (Suc n) \<longleftrightarrow> bit (Int.Pos m) n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3240
  \<open>bit (Int.Pos (num.Bit1 m)) (Suc n) \<longleftrightarrow> bit (Int.Pos m) n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3241
  \<open>bit (Int.Neg (num.Bit0 m)) (Suc n) \<longleftrightarrow> bit (Int.Neg m) n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3242
  \<open>bit (Int.Neg (num.Bit1 m)) (Suc n) \<longleftrightarrow> bit (Int.Neg (Num.inc m)) n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3243
  by (simp_all add: Num.add_One bit_0 bit_Suc)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3244
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3245
lemma not_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3246
  \<open>NOT (0 :: int) = - 1\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3247
  \<open>NOT (Int.Pos n) = Int.Neg (Num.inc n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3248
  \<open>NOT (Int.Neg n) = Num.sub n num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3249
  by (simp_all add: Num.add_One not_int_def)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3250
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3251
fun and_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3252
where
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3253
  \<open>and_num num.One num.One = Some num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3254
| \<open>and_num num.One (num.Bit0 n) = None\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3255
| \<open>and_num num.One (num.Bit1 n) = Some num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3256
| \<open>and_num (num.Bit0 m) num.One = None\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3257
| \<open>and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3258
| \<open>and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3259
| \<open>and_num (num.Bit1 m) num.One = Some num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3260
| \<open>and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3261
| \<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>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3262
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3263
context linordered_euclidean_semiring_bit_operations
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3264
begin
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3265
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3266
lemma numeral_and_num:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3267
  \<open>numeral m AND numeral n = (case and_num m n of None \<Rightarrow> 0 | Some n' \<Rightarrow> numeral n')\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3268
  by (induction m n rule: and_num.induct) (simp_all add: split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3269
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3270
lemma and_num_eq_None_iff:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3271
  \<open>and_num m n = None \<longleftrightarrow> numeral m AND numeral n = 0\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3272
  by (simp add: numeral_and_num split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3273
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3274
lemma and_num_eq_Some_iff:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3275
  \<open>and_num m n = Some q \<longleftrightarrow> numeral m AND numeral n = numeral q\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3276
  by (simp add: numeral_and_num split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3277
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3278
end
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3279
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3280
lemma and_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3281
  fixes i j :: int shows
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3282
  \<open>0 AND j = 0\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3283
  \<open>i AND 0 = 0\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3284
  \<open>Int.Pos n AND Int.Pos m = (case and_num n m of None \<Rightarrow> 0 | Some n' \<Rightarrow> Int.Pos n')\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3285
  \<open>Int.Neg n AND Int.Neg m = NOT (Num.sub n num.One OR Num.sub m num.One)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3286
  \<open>Int.Pos n AND Int.Neg num.One = Int.Pos n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3287
  \<open>Int.Pos n AND Int.Neg (num.Bit0 m) = Num.sub (or_not_num_neg (Num.BitM m) n) num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3288
  \<open>Int.Pos n AND Int.Neg (num.Bit1 m) = Num.sub (or_not_num_neg (num.Bit0 m) n) num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3289
  \<open>Int.Neg num.One AND Int.Pos m = Int.Pos m\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3290
  \<open>Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (or_not_num_neg (Num.BitM n) m) num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3291
  \<open>Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (or_not_num_neg (num.Bit0 n) m) num.One\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3292
  apply (auto simp: and_num_eq_None_iff [where ?'a = int] and_num_eq_Some_iff [where ?'a = int]
79069
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3293
    split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3294
     apply (simp_all only: sub_one_eq_not_neg numeral_or_not_num_eq minus_minus and_not_numerals
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3295
       bit.de_Morgan_disj bit.double_compl and_not_num_eq_None_iff and_not_num_eq_Some_iff ac_simps)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3296
  done
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3297
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3298
context linordered_euclidean_semiring_bit_operations
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3299
begin
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3300
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3301
fun or_num :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3302
where
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3303
  \<open>or_num num.One num.One = num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3304
| \<open>or_num num.One (num.Bit0 n) = num.Bit1 n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3305
| \<open>or_num num.One (num.Bit1 n) = num.Bit1 n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3306
| \<open>or_num (num.Bit0 m) num.One = num.Bit1 m\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3307
| \<open>or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3308
| \<open>or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3309
| \<open>or_num (num.Bit1 m) num.One = num.Bit1 m\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3310
| \<open>or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3311
| \<open>or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3312
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3313
lemma numeral_or_num:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3314
  \<open>numeral m OR numeral n = numeral (or_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3315
  by (induction m n rule: or_num.induct) simp_all
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3316
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3317
lemma numeral_or_num_eq:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3318
  \<open>numeral (or_num m n) = numeral m OR numeral n\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3319
  by (simp add: numeral_or_num)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3320
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3321
end
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3322
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3323
lemma or_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3324
  fixes i j :: int shows
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3325
  \<open>0 OR j = j\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3326
  \<open>i OR 0 = i\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3327
  \<open>Int.Pos n OR Int.Pos m = Int.Pos (or_num n m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3328
  \<open>Int.Neg n OR Int.Neg m = NOT (Num.sub n num.One AND Num.sub m num.One)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3329
  \<open>Int.Pos n OR Int.Neg num.One = Int.Neg num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3330
  \<open>Int.Pos n OR Int.Neg (num.Bit0 m) = (case and_not_num (Num.BitM m) n of None \<Rightarrow> -1 | Some n' \<Rightarrow> Int.Neg (Num.inc n'))\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3331
  \<open>Int.Pos n OR Int.Neg (num.Bit1 m) = (case and_not_num (num.Bit0 m) n of None \<Rightarrow> -1 | Some n' \<Rightarrow> Int.Neg (Num.inc n'))\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3332
  \<open>Int.Neg num.One OR Int.Pos m = Int.Neg num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3333
  \<open>Int.Neg (num.Bit0 n) OR Int.Pos m = (case and_not_num (Num.BitM n) m of None \<Rightarrow> -1 | Some n' \<Rightarrow> Int.Neg (Num.inc n'))\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3334
  \<open>Int.Neg (num.Bit1 n) OR Int.Pos m = (case and_not_num (num.Bit0 n) m of None \<Rightarrow> -1 | Some n' \<Rightarrow> Int.Neg (Num.inc n'))\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3335
  apply (auto simp: numeral_or_num_eq split: option.splits)
79069
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3336
         apply (simp_all only: and_not_num_eq_None_iff and_not_num_eq_Some_iff and_not_numerals
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3337
           numeral_or_not_num_eq or_eq_not_not_and bit.double_compl ac_simps flip: numeral_eq_iff [where ?'a = int])
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3338
         apply simp_all
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3339
  done
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3340
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3341
fun xor_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3342
where
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3343
  \<open>xor_num num.One num.One = None\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3344
| \<open>xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3345
| \<open>xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3346
| \<open>xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3347
| \<open>xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3348
| \<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>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3349
| \<open>xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3350
| \<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>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3351
| \<open>xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3352
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3353
context linordered_euclidean_semiring_bit_operations
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3354
begin
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3355
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3356
lemma numeral_xor_num:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3357
  \<open>numeral m XOR numeral n = (case xor_num m n of None \<Rightarrow> 0 | Some n' \<Rightarrow> numeral n')\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3358
  by (induction m n rule: xor_num.induct) (simp_all split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3359
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3360
lemma xor_num_eq_None_iff:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3361
  \<open>xor_num m n = None \<longleftrightarrow> numeral m XOR numeral n = 0\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3362
  by (simp add: numeral_xor_num split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3363
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3364
lemma xor_num_eq_Some_iff:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3365
  \<open>xor_num m n = Some q \<longleftrightarrow> numeral m XOR numeral n = numeral q\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3366
  by (simp add: numeral_xor_num split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3367
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3368
end
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3369
81876
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3370
context semiring_bit_operations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3371
begin
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3372
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3373
lemma push_bit_eq_pow:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3374
  \<open>push_bit (numeral n) 1 = numeral (Num.pow (Num.Bit0 Num.One) n)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3375
  by simp
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3376
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3377
lemma set_bit_of_0 [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3378
  \<open>set_bit n 0 = 2 ^ n\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3379
  by (simp add: set_bit_eq_or)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3380
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3381
lemma unset_bit_of_0 [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3382
  \<open>unset_bit n 0 = 0\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3383
  by (simp add: unset_bit_eq_or_xor)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3384
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3385
lemma flip_bit_of_0 [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3386
  \<open>flip_bit n 0 = 2 ^ n\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3387
  by (simp add: flip_bit_eq_xor)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3388
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3389
lemma set_bit_0_numeral_eq [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3390
  \<open>set_bit 0 (numeral Num.One) = 1\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3391
  \<open>set_bit 0 (numeral (Num.Bit0 m)) = numeral (Num.Bit1 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3392
  \<open>set_bit 0 (numeral (Num.Bit1 m)) = numeral (Num.Bit1 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3393
  by (simp_all add: set_bit_0)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3394
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3395
lemma set_bit_numeral_eq_or [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3396
  \<open>set_bit (numeral n) (numeral m) = numeral m OR push_bit (numeral n) 1\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3397
  by (fact set_bit_eq_or)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3398
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3399
lemma unset_bit_0_numeral_eq_and_not' [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3400
  \<open>unset_bit 0 (numeral Num.One) = 0\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3401
  \<open>unset_bit 0 (numeral (Num.Bit0 m)) = numeral (Num.Bit0 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3402
  \<open>unset_bit 0 (numeral (Num.Bit1 m)) = numeral (Num.Bit0 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3403
  by (simp_all add: unset_bit_0)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3404
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3405
lemma unset_bit_numeral_eq_or [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3406
  \<open>unset_bit (numeral n) (numeral m) =
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3407
    (case and_not_num m (Num.pow (Num.Bit0 Num.One) n)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3408
     of None \<Rightarrow> 0
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3409
      | Some q \<Rightarrow> numeral q)\<close> (is \<open>?lhs = _\<close>)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3410
proof -
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3411
  have \<open>?lhs = of_nat (unset_bit (numeral n) (numeral m))\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3412
    by (simp add: of_nat_unset_bit_eq)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3413
  also have \<open>unset_bit (numeral n) (numeral m) = nat (unset_bit (numeral n) (numeral m))\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3414
    by (simp flip: int_int_eq add: Bit_Operations.of_nat_unset_bit_eq)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3415
  finally have *: \<open>?lhs = of_nat (nat (unset_bit (numeral n) (numeral m)))\<close> .
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3416
  show ?thesis
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3417
    by (simp only: * unset_bit_eq_and_not Bit_Operations.push_bit_eq_pow int_numeral_and_not_num)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3418
      (auto split: option.splits)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3419
qed
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3420
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3421
lemma flip_bit_0_numeral_eq_or [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3422
  \<open>flip_bit 0 (numeral Num.One) = 0\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3423
  \<open>flip_bit 0 (numeral (Num.Bit0 m)) = numeral (Num.Bit1 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3424
  \<open>flip_bit 0 (numeral (Num.Bit1 m)) = numeral (Num.Bit0 m)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3425
  by (simp_all add: flip_bit_0)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3426
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3427
lemma flip_bit_numeral_eq_xor [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3428
  \<open>flip_bit (numeral n) (numeral m) = numeral m XOR push_bit (numeral n) 1\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3429
  by (fact flip_bit_eq_xor)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3430
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3431
end
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3432
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3433
context ring_bit_operations
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3434
begin
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3435
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3436
lemma set_bit_minus_numeral_eq_or [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3437
  \<open>set_bit (numeral n) (- numeral m) = - numeral m OR push_bit (numeral n) 1\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3438
  by (fact set_bit_eq_or)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3439
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3440
lemma unset_bit_minus_numeral_eq_and_not [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3441
  \<open>unset_bit (numeral n) (- numeral m) = - numeral m AND NOT (push_bit (numeral n) 1)\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3442
  by (fact unset_bit_eq_and_not)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3443
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3444
lemma flip_bit_minus_numeral_eq_xor [simp]:
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3445
  \<open>flip_bit (numeral n) (- numeral m) = - numeral m XOR push_bit (numeral n) 1\<close>
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3446
  by (fact flip_bit_eq_xor)
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3447
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3448
end
ac0716ca151b systematic checks for bit operations and more rules on symbolic terms
haftmann
parents: 81722
diff changeset
  3449
79069
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3450
lemma xor_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3451
  fixes i j :: int shows
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3452
  \<open>0 XOR j = j\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3453
  \<open>i XOR 0 = i\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3454
  \<open>Int.Pos n XOR Int.Pos m = (case xor_num n m of None \<Rightarrow> 0 | Some n' \<Rightarrow> Int.Pos n')\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3455
  \<open>Int.Neg n XOR Int.Neg m = Num.sub n num.One XOR Num.sub m num.One\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3456
  \<open>Int.Neg n XOR Int.Pos m = NOT (Num.sub n num.One XOR Int.Pos m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3457
  \<open>Int.Pos n XOR Int.Neg m = NOT (Int.Pos n XOR Num.sub m num.One)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3458
  by (simp_all add: xor_num_eq_None_iff [where ?'a = int] xor_num_eq_Some_iff [where ?'a = int] split: option.split)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3459
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3460
lemma push_bit_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3461
  \<open>push_bit 0 i = i\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3462
  \<open>push_bit (Suc n) i = push_bit n (Int.dup i)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3463
  by (simp_all add: ac_simps)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3464
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3465
lemma drop_bit_int_code [code]:
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3466
  fixes i :: int shows
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3467
  \<open>drop_bit 0 i = i\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3468
  \<open>drop_bit (Suc n) 0 = (0 :: int)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3469
  \<open>drop_bit (Suc n) (Int.Pos num.One) = 0\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3470
  \<open>drop_bit (Suc n) (Int.Pos (num.Bit0 m)) = drop_bit n (Int.Pos m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3471
  \<open>drop_bit (Suc n) (Int.Pos (num.Bit1 m)) = drop_bit n (Int.Pos m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3472
  \<open>drop_bit (Suc n) (Int.Neg num.One) = - 1\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3473
  \<open>drop_bit (Suc n) (Int.Neg (num.Bit0 m)) = drop_bit n (Int.Neg m)\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3474
  \<open>drop_bit (Suc n) (Int.Neg (num.Bit1 m)) = drop_bit n (Int.Neg (Num.inc m))\<close>
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3475
  by (simp_all add: drop_bit_Suc add_One)
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3476
48ca09068adf grouped lemmas for symbolic computations
haftmann
parents: 79068
diff changeset
  3477
74108
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  3478
subsection \<open>More properties\<close>
3146646a43a7 simplified hierarchy of type classes for bit operations
haftmann
parents: 74101
diff changeset
  3479
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3480
lemma take_bit_eq_mask_iff:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3481
  \<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
  3482
  for k :: int
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3483
proof
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3484
  assume ?P
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3485
  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
  3486
    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
  3487
  then show ?Q
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3488
    by (simp only: take_bit_add)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3489
next
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3490
  assume ?Q
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3491
  then have \<open>take_bit n (k + 1) - 1 = - 1\<close>
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3492
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3493
  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
  3494
    by simp
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3495
  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
  3496
    by (simp add: take_bit_eq_mod mod_simps)
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3497
  ultimately show ?P
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3498
    by simp
72830
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3499
qed
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3500
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3501
lemma take_bit_eq_mask_iff_exp_dvd:
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3502
  \<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
  3503
  for k :: int
ec0d3a62bb3b moved some lemmas from AFP to distribution
haftmann
parents: 72792
diff changeset
  3504
  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
  3505
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  3506
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3507
subsection \<open>Bit concatenation\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3508
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3509
definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close>
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3510
  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
  3511
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  3512
lemma bit_concat_bit_iff [bit_simps]:
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3513
  \<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
  3514
  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
  3515
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3516
lemma concat_bit_eq:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3517
  \<open>concat_bit n k l = take_bit n k + push_bit n l\<close>
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3518
proof -
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3519
  have \<open>take_bit n k AND push_bit n l = 0\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3520
    by (simp add: bit_eq_iff bit_simps)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3521
  then show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3522
    by (simp add: bit_eq_iff bit_simps disjunctive_add_eq_or)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3523
qed
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3524
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3525
lemma concat_bit_0 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3526
  \<open>concat_bit 0 k l = l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3527
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3528
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3529
lemma concat_bit_Suc:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3530
  \<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
  3531
  by (simp add: concat_bit_eq take_bit_Suc push_bit_double)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3532
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3533
lemma concat_bit_of_zero_1 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3534
  \<open>concat_bit n 0 l = push_bit n l\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3535
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3536
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3537
lemma concat_bit_of_zero_2 [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3538
  \<open>concat_bit n k 0 = take_bit n k\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3539
  by (simp add: concat_bit_def take_bit_eq_mask)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3540
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3541
lemma concat_bit_nonnegative_iff [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3542
  \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3543
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3544
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3545
lemma concat_bit_negative_iff [simp]:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3546
  \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close>
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3547
  by (simp add: concat_bit_def)
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3548
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3549
lemma concat_bit_assoc:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3550
  \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3551
  by (rule bit_eqI) (auto simp: bit_concat_bit_iff ac_simps)
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3552
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3553
lemma concat_bit_assoc_sym:
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3554
  \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3555
  by (rule bit_eqI) (auto simp: bit_concat_bit_iff ac_simps min_def)
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3556
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3557
lemma concat_bit_eq_iff:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3558
  \<open>concat_bit n k l = concat_bit n r s
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3559
    \<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
  3560
proof
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3561
  assume ?Q
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3562
  then show ?P
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3563
    by (simp add: concat_bit_def)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3564
next
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3565
  assume ?P
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3566
  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
  3567
    by (simp add: bit_eq_iff)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3568
  have \<open>take_bit n k = take_bit n r\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3569
  proof (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3570
    fix m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3571
    from * [of m]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3572
    show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3573
      by (auto simp: bit_take_bit_iff bit_concat_bit_iff)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3574
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3575
  moreover have \<open>push_bit n l = push_bit n s\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3576
  proof (rule bit_eqI)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3577
    fix m
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3578
    from * [of m]
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3579
    show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3580
      by (auto simp: bit_push_bit_iff bit_concat_bit_iff)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3581
  qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3582
  then have \<open>l = s\<close>
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3583
    by (simp add: push_bit_eq_mult)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3584
  ultimately show ?Q
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3585
    by (simp add: concat_bit_def)
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3586
qed
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3587
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3588
lemma take_bit_concat_bit_eq:
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3589
  \<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
  3590
  by (rule bit_eqI)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3591
    (auto simp: bit_take_bit_iff bit_concat_bit_iff min_def)
72227
0f3d24dc197f more on conversions
haftmann
parents: 72187
diff changeset
  3592
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  3593
lemma concat_bit_take_bit_eq:
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  3594
  \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  3595
  by (simp add: concat_bit_def [abs_def])
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 72397
diff changeset
  3596
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3597
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3598
subsection \<open>Taking bits with sign propagation\<close>
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3599
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3600
context ring_bit_operations
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3601
begin
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3602
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3603
definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3604
  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
  3605
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3606
lemma signed_take_bit_eq_if_positive:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3607
  \<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
  3608
  using that by (simp add: signed_take_bit_def)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3609
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3610
lemma signed_take_bit_eq_if_negative:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3611
  \<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
  3612
  using that by (simp add: signed_take_bit_def)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3613
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3614
lemma even_signed_take_bit_iff:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3615
  \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3616
  by (auto simp: bit_0 signed_take_bit_def even_or_iff even_mask_iff bit_double_iff)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3617
72611
c7bc3e70a8c7 official collection for bit projection simplifications
haftmann
parents: 72512
diff changeset
  3618
lemma bit_signed_take_bit_iff [bit_simps]:
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  3619
  \<open>bit (signed_take_bit m a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit a (min m n)\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3620
  by (simp add: signed_take_bit_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff min_def not_le)
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  3621
    (blast dest: bit_imp_possible_bit)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3622
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3623
lemma signed_take_bit_0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3624
  \<open>signed_take_bit 0 a = - (a mod 2)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  3625
  by (simp add: bit_0 signed_take_bit_def odd_iff_mod_2_eq_one)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3626
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3627
lemma signed_take_bit_Suc:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3628
  \<open>signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\<close>
75085
ccc3a72210e6 Avoid overaggresive simplification.
haftmann
parents: 74618
diff changeset
  3629
  by (simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 possible_bit_less_imp flip: bit_Suc min_Suc_Suc)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3630
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3631
lemma signed_take_bit_of_0 [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3632
  \<open>signed_take_bit n 0 = 0\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3633
  by (simp add: signed_take_bit_def)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3634
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3635
lemma signed_take_bit_of_minus_1 [simp]:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3636
  \<open>signed_take_bit n (- 1) = - 1\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3637
  by (simp add: signed_take_bit_def mask_eq_exp_minus_1 possible_bit_def)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3638
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3639
lemma signed_take_bit_Suc_1 [simp]:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3640
  \<open>signed_take_bit (Suc n) 1 = 1\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3641
  by (simp add: signed_take_bit_Suc)
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3642
74497
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  3643
lemma signed_take_bit_numeral_of_1 [simp]:
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  3644
  \<open>signed_take_bit (numeral k) 1 = 1\<close>
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  3645
  by (simp add: bit_1_iff signed_take_bit_eq_if_positive)
9c04a82c3128 more complete simp rules
haftmann
parents: 74495
diff changeset
  3646
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3647
lemma signed_take_bit_rec:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3648
  \<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
  3649
  by (cases n) (simp_all add: signed_take_bit_Suc)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3650
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3651
lemma signed_take_bit_eq_iff_take_bit_eq:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3652
  \<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
  3653
proof -
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3654
  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
  3655
    by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def)
74309
42523fbf643b explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents: 74163
diff changeset
  3656
      (use bit_imp_possible_bit in fastforce)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3657
  then show ?thesis
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3658
    by (auto simp: fun_eq_iff intro: bit_eqI)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3659
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3660
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3661
lemma signed_take_bit_signed_take_bit [simp]:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3662
  \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3663
  by (auto simp: bit_eq_iff bit_simps ac_simps)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3664
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3665
lemma signed_take_bit_take_bit:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3666
  \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3667
  by (rule bit_eqI) (auto simp: bit_signed_take_bit_iff min_def bit_take_bit_iff)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3668
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3669
lemma take_bit_signed_take_bit:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3670
  \<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
  3671
  using that by (rule le_SucE; intro bit_eqI)
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3672
   (auto simp: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq)
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3673
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3674
lemma signed_take_bit_eq_take_bit_add:
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3675
  \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n)) * of_bool (bit k n)\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3676
proof (cases \<open>bit k n\<close>)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3677
  case False
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3678
  show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3679
    by (rule bit_eqI) (simp add: False bit_simps min_def less_Suc_eq)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3680
next
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3681
  case True
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3682
  have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3683
    by (rule bit_eqI) (auto simp: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3684
  also have \<open>\<dots> = take_bit (Suc n) k + NOT (mask (Suc n))\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3685
    by (simp add: disjunctive_add_eq_or bit_eq_iff bit_simps)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3686
  finally show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3687
    by (simp add: True)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3688
qed
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3689
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3690
lemma signed_take_bit_eq_take_bit_minus:
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3691
  \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3692
  by (simp add: signed_take_bit_eq_take_bit_add flip: minus_exp_eq_not_mask)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3693
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3694
end
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3695
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3696
text \<open>Modulus centered around 0\<close>
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3697
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3698
lemma signed_take_bit_eq_concat_bit:
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3699
  \<open>signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\<close>
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74498
diff changeset
  3700
  by (simp add: concat_bit_def signed_take_bit_def)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3701
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3702
lemma signed_take_bit_add:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3703
  \<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
  3704
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3705
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3706
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3707
     (take_bit (Suc n) (signed_take_bit n k) +
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3708
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3709
    take_bit (Suc n) (k + l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3710
    by (simp add: take_bit_signed_take_bit take_bit_add)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3711
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3712
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3713
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3714
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3715
lemma signed_take_bit_diff:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3716
  \<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
  3717
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3718
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3719
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3720
     (take_bit (Suc n) (signed_take_bit n k) -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3721
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3722
    take_bit (Suc n) (k - l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3723
    by (simp add: take_bit_signed_take_bit take_bit_diff)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3724
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3725
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3726
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3727
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3728
lemma signed_take_bit_minus:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3729
  \<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
  3730
  for k :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3731
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3732
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3733
     (- take_bit (Suc n) (signed_take_bit n k)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3734
    take_bit (Suc n) (- k)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3735
    by (simp add: take_bit_signed_take_bit take_bit_minus)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3736
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3737
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3738
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3739
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3740
lemma signed_take_bit_mult:
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3741
  \<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
  3742
  for k l :: int
72187
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3743
proof -
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3744
  have \<open>take_bit (Suc n)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3745
     (take_bit (Suc n) (signed_take_bit n k) *
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3746
      take_bit (Suc n) (signed_take_bit n l)) =
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3747
    take_bit (Suc n) (k * l)\<close>
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3748
    by (simp add: take_bit_signed_take_bit take_bit_mult)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3749
  then show ?thesis
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3750
    by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult)
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3751
qed
e4aecb0c7296 more lemmas
haftmann
parents: 72130
diff changeset
  3752
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3753
lemma signed_take_bit_eq_take_bit_shift:
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3754
  \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close> (is \<open>?lhs = ?rhs\<close>)
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3755
  for k :: int
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3756
proof -
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3757
  have \<open>take_bit n k AND 2 ^ n = 0\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3758
    by (rule bit_eqI) (simp add: bit_simps)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3759
  then have *: \<open>take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3760
    by (simp add: disjunctive_add_eq_or)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3761
  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
  3762
    by (simp add: minus_exp_eq_not_mask)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3763
  also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3764
    by (rule disjunctive_add_eq_or) (simp add: bit_eq_iff bit_simps)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3765
  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
  3766
  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
  3767
    by (simp only: take_bit_add)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3768
  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
  3769
    by (simp add: take_bit_Suc_from_most)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3770
  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
  3771
    by (simp add: ac_simps)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3772
  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>
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3773
    by (rule disjunctive_add_eq_or, rule bit_eqI) (simp add: bit_simps)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3774
  finally show ?thesis
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3775
    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
  3776
qed
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3777
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3778
lemma signed_take_bit_nonnegative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3779
  \<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
  3780
  for k :: int
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3781
  by (simp add: signed_take_bit_def not_less concat_bit_def)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3782
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3783
lemma signed_take_bit_negative_iff [simp]:
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3784
  \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3785
  for k :: int
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3786
  by (simp add: signed_take_bit_def not_less concat_bit_def)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3787
73868
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3788
lemma signed_take_bit_int_greater_eq_minus_exp [simp]:
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3789
  \<open>- (2 ^ n) \<le> signed_take_bit n k\<close>
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3790
  for k :: int
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3791
  by (simp add: signed_take_bit_eq_take_bit_shift)
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3792
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3793
lemma signed_take_bit_int_less_exp [simp]:
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3794
  \<open>signed_take_bit n k < 2 ^ n\<close>
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3795
  for k :: int
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3796
  using take_bit_int_less_exp [of \<open>Suc n\<close>]
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3797
  by (simp add: signed_take_bit_eq_take_bit_shift)
465846b611d5 some word streamlining
haftmann
parents: 73816
diff changeset
  3798
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3799
lemma signed_take_bit_int_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3800
  \<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
  3801
  for k :: int
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3802
  by (auto simp: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps)
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3803
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3804
lemma signed_take_bit_int_eq_self:
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3805
  \<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
  3806
  for k :: int
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3807
  using that by (simp add: signed_take_bit_int_eq_self_iff)
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3808
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3809
lemma signed_take_bit_int_less_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3810
  \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3811
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3812
  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
  3813
    linarith
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3814
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3815
lemma signed_take_bit_int_less_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3816
  \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3817
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3818
  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
  3819
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3820
lemma signed_take_bit_int_greater_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3821
  \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3822
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3823
  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
  3824
    linarith
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3825
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3826
lemma signed_take_bit_int_greater_eq_self_iff:
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3827
  \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3828
  for k :: int
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3829
  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
  3830
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3831
lemma signed_take_bit_int_greater_eq:
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3832
  \<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
  3833
  for k :: int
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3834
  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
  3835
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3836
72261
5193570b739a more lemmas
haftmann
parents: 72241
diff changeset
  3837
lemma signed_take_bit_int_less_eq:
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3838
  \<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
  3839
  for k :: int
72262
a282abb07642 integrated generic conversions into word corpse
haftmann
parents: 72261
diff changeset
  3840
  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
  3841
  by (simp add: signed_take_bit_eq_take_bit_shift)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3842
82524
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
  3843
lemma signed_take_bit_Suc_sgn_eq [simp]:
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
  3844
  \<open>signed_take_bit (Suc n) (sgn k) = sgn k\<close> for k :: int
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
  3845
  by (simp add: sgn_if)
df5b2785abd6 more lemmas
haftmann
parents: 82289
diff changeset
  3846
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3847
lemma signed_take_bit_Suc_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3848
  \<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
  3849
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3850
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3851
lemma signed_take_bit_Suc_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3852
  \<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
  3853
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3854
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3855
lemma signed_take_bit_Suc_minus_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3856
  \<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
  3857
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3858
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3859
lemma signed_take_bit_Suc_minus_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3860
  \<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
  3861
  by (simp add: signed_take_bit_Suc)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3862
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3863
lemma signed_take_bit_numeral_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3864
  \<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
  3865
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3866
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3867
lemma signed_take_bit_numeral_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3868
  \<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
  3869
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3870
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3871
lemma signed_take_bit_numeral_minus_bit0 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3872
  \<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
  3873
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3874
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3875
lemma signed_take_bit_numeral_minus_bit1 [simp]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3876
  \<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
  3877
  by (simp add: signed_take_bit_rec)
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3878
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3879
lemma signed_take_bit_code [code]:
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3880
  \<open>signed_take_bit n a =
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3881
  (let l = take_bit (Suc n) a
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3882
   in if bit l n then l + push_bit (Suc n) (- 1) else l)\<close>
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  3883
  by (simp add: signed_take_bit_eq_take_bit_add bit_simps)
72010
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3884
a851ce626b78 signed_take_bit
haftmann
parents: 72009
diff changeset
  3885
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3886
subsection \<open>Key ideas of bit operations\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3887
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3888
text \<open>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3889
  When formalizing bit operations, it is tempting to represent
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3890
  bit values as explicit lists over a binary type. This however
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3891
  is a bad idea, mainly due to the inherent ambiguities in
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3892
  representation concerning repeating leading bits.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3893
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3894
  Hence this approach avoids such explicit lists altogether
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3895
  following an algebraic path:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3896
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3897
  \<^item> Bit values are represented by numeric types: idealized
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3898
    unbounded bit values can be represented by type \<^typ>\<open>int\<close>,
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3899
    bounded bit values by quotient types over \<^typ>\<open>int\<close>.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3900
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3901
  \<^item> (A special case are idealized unbounded bit values ending
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3902
    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
  3903
    only support a restricted set of operations).
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3904
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3905
  \<^item> From this idea follows that
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3906
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3907
      \<^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
  3908
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3909
      \<^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
  3910
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3911
  \<^item> Concerning bounded bit values, iterated shifts to the left
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3912
    may result in eliminating all bits by shifting them all
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3913
    beyond the boundary.  The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3914
    represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary.
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3915
71965
d45f5d4c41bd more class operations for the sake of efficient generated code
haftmann
parents: 71956
diff changeset
  3916
  \<^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
  3917
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3918
  \<^item> This leads to the most fundamental properties of bit values:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3919
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3920
      \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3921
79480
c7cb1bf6efa0 consolidated name of lemma analogously to nat/int/word_bit_induct
haftmann
parents: 79117
diff changeset
  3922
      \<^item> Induction rule: @{thm bit_induct [where ?'a = int, no_vars]}
71800
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3923
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3924
  \<^item> Typical operations are characterized as follows:
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3925
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3926
      \<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3927
71956
a4bffc0de967 bit operations as distinctive library theory
haftmann
parents: 71922
diff changeset
  3928
      \<^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
  3929
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3930
      \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3931
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3932
      \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3933
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3934
      \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3935
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3936
      \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3937
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3938
      \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3939
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3940
      \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3941
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3942
      \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3943
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3944
      \<^item> Set a single bit: @{thm set_bit_eq_or [where ?'a = int, no_vars]}
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3945
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3946
      \<^item> Unset a single bit: @{thm unset_bit_eq_and_not [where ?'a = int, no_vars]}
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3947
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3948
      \<^item> Flip a single bit: @{thm flip_bit_eq_xor [where ?'a = int, no_vars]}
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3949
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3950
      \<^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
  3951
72241
5a6d8675bf4b generalized signed_take_bit
haftmann
parents: 72239
diff changeset
  3952
      \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]}
72028
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3953
08f1e4cb735f concatentation of bit values
haftmann
parents: 72023
diff changeset
  3954
      \<^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
  3955
\<close>
35a951ed2e82 documentation of relevant ideas
haftmann
parents: 71535
diff changeset
  3956
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3957
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3958
subsection \<open>Lemma duplicates and other\<close>
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  3959
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3960
context semiring_bits
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3961
begin
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3962
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  3963
lemma exp_div_exp_eq:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3964
  \<open>2 ^ m div 2 ^ n = of_bool (2 ^ m \<noteq> 0 \<and> m \<ge> n) * 2 ^ (m - n)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3965
  using bit_exp_iff div_exp_eq
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  3966
  by (intro bit_eqI) (auto simp: bit_iff_odd possible_bit_def)
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3967
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  3968
lemma bits_1_div_2:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3969
  \<open>1 div 2 = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3970
  by (fact half_1)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3971
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  3972
lemma bits_1_div_exp:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3973
  \<open>1 div 2 ^ n = of_bool (n = 0)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3974
  using div_exp_eq [of 1 1] by (cases n) simp_all
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3975
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  3976
lemma exp_add_not_zero_imp:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3977
  \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> if \<open>2 ^ (m + n) \<noteq> 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3978
proof -
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3979
  have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3980
  proof (rule notI)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3981
    assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3982
    then have \<open>2 ^ (m + n) = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3983
      by (rule disjE) (simp_all add: power_add)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3984
    with that show False ..
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3985
  qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3986
  then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3987
    by simp_all
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3988
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3989
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3990
lemma
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  3991
  exp_add_not_zero_imp_left: \<open>2 ^ m \<noteq> 0\<close>
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  3992
  and exp_add_not_zero_imp_right: \<open>2 ^ n \<noteq> 0\<close>
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3993
  if \<open>2 ^ (m + n) \<noteq> 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3994
proof -
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3995
  have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3996
  proof (rule notI)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3997
    assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3998
    then have \<open>2 ^ (m + n) = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  3999
      by (rule disjE) (simp_all add: power_add)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4000
    with that show False ..
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4001
  qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4002
  then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4003
    by simp_all
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4004
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4005
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4006
lemma exp_not_zero_imp_exp_diff_not_zero:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4007
  \<open>2 ^ (n - m) \<noteq> 0\<close> if \<open>2 ^ n \<noteq> 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4008
proof (cases \<open>m \<le> n\<close>)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4009
  case True
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4010
  moreover define q where \<open>q = n - m\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4011
  ultimately have \<open>n = m + q\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4012
    by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4013
  with that show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4014
    by (simp add: exp_add_not_zero_imp_right)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4015
next
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4016
  case False
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4017
  with that show ?thesis
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4018
    by simp
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4019
qed
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4020
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4021
lemma exp_eq_0_imp_not_bit:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4022
  \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4023
  using that by (simp add: bit_iff_odd)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4024
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4025
lemma bit_disjunctive_add_iff:
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4026
  \<open>bit (a + b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4027
  if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4028
proof (cases \<open>possible_bit TYPE('a) n\<close>)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4029
  case False
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4030
  then show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4031
    by (auto dest: impossible_bit)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4032
next
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4033
  case True
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4034
  with that show ?thesis proof (induction n arbitrary: a b)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4035
    case 0
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4036
    from "0.prems"(1) [of 0] show ?case
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  4037
      by (auto simp: bit_0)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4038
  next
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4039
    case (Suc n)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4040
    from Suc.prems(1) [of 0] have even: \<open>even a \<or> even b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  4041
      by (auto simp: bit_0)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4042
    have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4043
      using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4044
    from Suc.prems(2) have \<open>possible_bit TYPE('a) (Suc n)\<close> \<open>possible_bit TYPE('a) n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4045
      by (simp_all add: possible_bit_less_imp)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4046
    have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4047
      using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4048
    also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  4049
      using even by (auto simp: algebra_simps mod2_eq_if)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4050
    finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4051
      using \<open>possible_bit TYPE('a) (Suc n)\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff possible_bit_def)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4052
    also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4053
      using bit \<open>possible_bit TYPE('a) n\<close> by (rule Suc.IH)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4054
    finally show ?case
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4055
      by (simp add: bit_Suc)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4056
  qed
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4057
qed
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4058
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4059
end
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4060
79116
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4061
context semiring_bit_operations
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4062
begin
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4063
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4064
lemma even_mask_div_iff:
79588
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
  4065
  \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  4066
  using bit_mask_iff [of m n] by (auto simp: mask_eq_exp_minus_1 bit_iff_odd possible_bit_def)
79588
9f22b71e209e simplified class specification
haftmann
parents: 79585
diff changeset
  4067
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4068
lemma mod_exp_eq:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4069
  \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4070
  by (simp flip: take_bit_eq_mod add: ac_simps)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4071
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4072
lemma mult_exp_mod_exp_eq:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4073
  \<open>m \<le> n \<Longrightarrow> (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4074
  by (simp flip: push_bit_eq_mult take_bit_eq_mod add: push_bit_take_bit)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4075
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4076
lemma div_exp_mod_exp_eq:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4077
  \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4078
  by (simp flip: drop_bit_eq_div take_bit_eq_mod add: drop_bit_take_bit)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4079
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4080
lemma even_mult_exp_div_exp_iff:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4081
  \<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>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4082
  by (simp flip: push_bit_eq_mult drop_bit_eq_div add: even_drop_bit_iff_not_bit bit_simps possible_bit_def) auto
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4083
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4084
lemma mod_exp_div_exp_eq_0:
79531
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4085
  \<open>a mod 2 ^ n div 2 ^ n = 0\<close>
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4086
  by (simp flip: take_bit_eq_mod drop_bit_eq_div add: drop_bit_take_bit)
22a137a6de44 rearranged and reformulated abstract classes for bit structures and operations
haftmann
parents: 79489
diff changeset
  4087
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4088
lemmas bits_one_mod_two_eq_one = one_mod_two_eq_one
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4089
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4090
lemmas set_bit_def = set_bit_eq_or
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4091
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4092
lemmas unset_bit_def = unset_bit_eq_and_not
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4093
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4094
lemmas flip_bit_def = flip_bit_eq_xor
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4095
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4096
lemma disjunctive_add:
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4097
  \<open>a + b = a OR b\<close> if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4098
  by (rule disjunctive_add_eq_or) (use that in \<open>simp add: bit_eq_iff bit_simps\<close>)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4099
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4100
lemma even_mod_exp_div_exp_iff:
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
  4101
  \<open>even (a mod 2 ^ m div 2 ^ n) \<longleftrightarrow> m \<le> n \<or> even (a div 2 ^ n)\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  4102
  by (auto simp: even_drop_bit_iff_not_bit bit_simps simp flip: drop_bit_eq_div take_bit_eq_mod)
79673
c172eecba85d simplified specification of type class semiring_bits
haftmann
parents: 79610
diff changeset
  4103
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4104
end
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4105
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4106
context ring_bit_operations
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4107
begin
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4108
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4109
lemma disjunctive_diff:
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4110
  \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4111
proof -
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4112
  have \<open>NOT a + b = NOT a OR b\<close>
80758
8f96e1329845 Some tidying
paulson <lp15@cam.ac.uk>
parents: 79893
diff changeset
  4113
    by (rule disjunctive_add) (auto simp: bit_not_iff dest: that)
79610
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4114
  then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4115
    by simp
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4116
  then show ?thesis
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4117
    by (simp add: not_add_distrib)
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4118
qed
ad29777e8746 more on disjunctive addition/subtraction
haftmann
parents: 79590
diff changeset
  4119
79116
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4120
end
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4121
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4122
lemma and_nat_rec:
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4123
  \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4124
  by (fact and_rec)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4125
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4126
lemma or_nat_rec:
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4127
  \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4128
  by (fact or_rec)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4129
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4130
lemma xor_nat_rec:
79070
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4131
  \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4132
  by (fact xor_rec)
a4775fe69f5d restructured
haftmann
parents: 79069
diff changeset
  4133
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4134
lemma bit_push_bit_iff_nat:
79071
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  4135
  \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  4136
  by (fact bit_push_bit_iff')
7ab8b3f1d84b generalized
haftmann
parents: 79070
diff changeset
  4137
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4138
lemma mask_half_int:
79116
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4139
  \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4140
  by (fact mask_half)
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4141
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4142
lemma not_int_rec:
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  4143
  \<open>NOT k = of_bool (even k) + 2 * NOT (k div 2)\<close> for k :: int
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  4144
  by (fact not_rec)
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  4145
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4146
lemma even_not_iff_int:
79068
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  4147
  \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  4148
  by (fact even_not_iff)
cb72e2c0c539 sorted out lemma duplicates
haftmann
parents: 79031
diff changeset
  4149
79072
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  4150
lemma bit_not_int_iff':
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  4151
  \<open>bit (- k - 1) n \<longleftrightarrow> \<not> bit k n\<close> for k :: int
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  4152
  by (simp flip: not_eq_complement add: bit_simps)
a91050cd5c93 de-duplicated specification of class ring_bit_operations
haftmann
parents: 79071
diff changeset
  4153
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4154
lemmas and_int_rec = and_int.rec
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4155
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4156
lemma even_and_iff_int:
79116
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4157
  \<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4158
  by (fact even_and_iff)
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4159
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4160
lemmas bit_and_int_iff = and_int.bit_iff
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4161
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4162
lemmas or_int_rec = or_int.rec
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4163
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4164
lemmas bit_or_int_iff = or_int.bit_iff
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4165
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4166
lemmas xor_int_rec = xor_int.rec
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4167
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4168
lemmas bit_xor_int_iff = xor_int.bit_iff
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4169
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4170
lemma drop_bit_push_bit_int:
79116
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4171
  \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4172
  by (fact drop_bit_push_bit)
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4173
79893
7ea70796acaa avoid [no_atp] declations shadowing propositions from sledgehammer
haftmann
parents: 79673
diff changeset
  4174
lemma bit_push_bit_iff_int:
79116
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4175
  \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4176
  by (fact bit_push_bit_iff')
b90bf6636260 explicit annotation of lemma duplicates
haftmann
parents: 79072
diff changeset
  4177
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  4178
bundle bit_operations_syntax
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 74097
diff changeset
  4179
begin
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  4180
notation
74391
930047942f46 repaired slip
haftmann
parents: 74364
diff changeset
  4181
  not  (\<open>NOT\<close>)
74364
99add5178e51 NOT is part of syntax bundle also
haftmann
parents: 74309
diff changeset
  4182
    and "and"  (infixr \<open>AND\<close> 64)
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  4183
    and or  (infixr \<open>OR\<close>  59)
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  4184
    and xor  (infixr \<open>XOR\<close> 59)
81132
dff7dfd8dce3 more robust declarations via "no syntax" bundles;
wenzelm
parents: 80758
diff changeset
  4185
end
dff7dfd8dce3 more robust declarations via "no syntax" bundles;
wenzelm
parents: 80758
diff changeset
  4186
dff7dfd8dce3 more robust declarations via "no syntax" bundles;
wenzelm
parents: 80758
diff changeset
  4187
unbundle no bit_operations_syntax
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 73969
diff changeset
  4188
71442
d45495e897f4 more instances
haftmann
parents: 71426
diff changeset
  4189
end