src/HOL/ex/Bit_Operations.thy
author haftmann
Mon, 03 Feb 2020 20:42:04 +0000
changeset 71418 bd9d27ccb3a3
parent 71413 65ffe9e910d4
child 71419 1d8e914e04d6
permissions -rw-r--r--
more theorems
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
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     4
section \<open>Proof of concept for purely algebraically founded lists of bits\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     5
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     6
theory Bit_Operations
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     7
  imports
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     8
    "HOL-Library.Boolean_Algebra"
71095
038727567817 tuned order between theories
haftmann
parents: 71094
diff changeset
     9
    Main
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    10
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    11
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    12
context semiring_bit_shifts
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    13
begin
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    14
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
    15
(*lemma bit_push_bit_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
    16
  \<open>bit (push_bit m a) n \<longleftrightarrow> n \<ge> m \<and> 2 ^ n \<noteq> 0 \<and> bit a (n - m)\<close>*)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
    17
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    18
end
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    19
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    20
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    21
subsection \<open>Bit operations in suitable algebraic structures\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    22
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    23
class semiring_bit_operations = semiring_bit_shifts +
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    24
  fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr "AND" 64)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    25
    and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr "OR"  59)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    26
    and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr "XOR" 59)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    27
  assumes bit_and_iff: \<open>\<And>n. bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    28
    and bit_or_iff: \<open>\<And>n. bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    29
    and bit_xor_iff: \<open>\<And>n. bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    30
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    31
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    32
text \<open>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    33
  We want the bitwise operations to bind slightly weaker
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71042
diff changeset
    34
  than \<open>+\<close> and \<open>-\<close>.
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    35
  For the sake of code generation
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    36
  the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    37
  are specified as definitional class operations.
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    38
\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    39
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    40
definition map_bit :: \<open>nat \<Rightarrow> (bool \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a\<close>
71181
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    41
  where \<open>map_bit n f a = take_bit n a + push_bit n (of_bool (f (bit a n)) + 2 * drop_bit (Suc n) a)\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    42
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    43
definition set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    44
  where \<open>set_bit n = map_bit n top\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    45
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    46
definition unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    47
  where \<open>unset_bit n = map_bit n bot\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    48
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    49
definition flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    50
  where \<open>flip_bit n = map_bit n Not\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    51
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    52
text \<open>
71181
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    53
  Having 
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    54
  \<^const>\<open>set_bit\<close>, \<^const>\<open>unset_bit\<close> and \<^const>\<open>flip_bit\<close> as separate
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    55
  operations allows to implement them using bit masks later.
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    56
\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    57
71181
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    58
lemma stable_imp_drop_eq:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    59
  \<open>drop_bit n a = a\<close> if \<open>a div 2 = a\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    60
  by (induction n) (simp_all add: that)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    61
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    62
lemma map_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    63
  \<open>map_bit 0 f a = of_bool (f (odd a)) + 2 * (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    64
  by (simp add: map_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    65
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    66
lemma map_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    67
  \<open>map_bit (Suc n) f a = a mod 2 + 2 * map_bit n f (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    68
  by (auto simp add: map_bit_def algebra_simps mod2_eq_if push_bit_add mult_2
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    69
    elim: evenE oddE)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    70
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    71
lemma set_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    72
  \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    73
  by (simp add: set_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    74
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    75
lemma set_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    76
  \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    77
  by (simp add: set_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    78
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    79
lemma unset_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    80
  \<open>unset_bit 0 a = 2 * (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    81
  by (simp add: unset_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    82
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    83
lemma unset_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    84
  \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    85
  by (simp add: unset_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    86
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    87
lemma flip_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    88
  \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    89
  by (simp add: flip_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    90
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    91
lemma flip_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    92
  \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    93
  by (simp add: flip_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    94
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    95
sublocale "and": semilattice \<open>(AND)\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    96
  by standard (auto simp add: bit_eq_iff bit_and_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    97
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    98
sublocale or: semilattice_neutr \<open>(OR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    99
  by standard (auto simp add: bit_eq_iff bit_or_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   100
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   101
sublocale xor: comm_monoid \<open>(XOR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   102
  by standard (auto simp add: bit_eq_iff bit_xor_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   103
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   104
lemma zero_and_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   105
  "0 AND a = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   106
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   107
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   108
lemma and_zero_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   109
  "a AND 0 = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   110
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   111
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   112
lemma one_and_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   113
  "1 AND a = of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   114
  by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   115
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   116
lemma and_one_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   117
  "a AND 1 = of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   118
  using one_and_eq [of a] by (simp add: ac_simps)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   119
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   120
lemma one_or_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   121
  "1 OR a = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   122
  by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   123
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   124
lemma or_one_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   125
  "a OR 1 = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   126
  using one_or_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   127
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   128
lemma one_xor_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   129
  "1 XOR a = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   130
  by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff odd_bit_iff_bit_pred elim: oddE)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   131
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   132
lemma xor_one_int_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   133
  "a XOR 1 = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   134
  using one_xor_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   135
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   136
lemma take_bit_and [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   137
  \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   138
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   139
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   140
lemma take_bit_or [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   141
  \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   142
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   143
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   144
lemma take_bit_xor [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   145
  \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   146
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   147
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   148
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   149
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   150
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
   151
  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   152
  assumes bit_not_iff: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   153
  assumes minus_eq_not_minus_1: \<open>- a = NOT (a - 1)\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   154
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   155
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   156
text \<open>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   157
  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
   158
  definitional class operation.  Note that \<^const>\<open>not\<close> has no
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   159
  sensible definition for unlimited but only positive bit strings
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   160
  (type \<^typ>\<open>nat\<close>).
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   161
\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   162
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   163
lemma bits_minus_1_mod_2_eq [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   164
  \<open>(- 1) mod 2 = 1\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   165
  by (simp add: mod_2_eq_odd)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   166
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   167
lemma not_eq_complement:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   168
  \<open>NOT a = - a - 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   169
  using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   170
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   171
lemma minus_eq_not_plus_1:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   172
  \<open>- a = NOT a + 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   173
  using not_eq_complement [of a] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   174
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   175
lemma bit_minus_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   176
  \<open>bit (- a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit (a - 1) n\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   177
  by (simp add: minus_eq_not_minus_1 bit_not_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   178
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   179
lemma even_not_iff [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   180
  "even (NOT a) \<longleftrightarrow> odd a"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   181
  using bit_not_iff [of a 0] by auto
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   182
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   183
lemma bit_not_exp_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   184
  \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<noteq> m\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   185
  by (auto simp add: bit_not_iff bit_exp_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   186
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   187
lemma bit_minus_1_iff [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   188
  \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   189
  by (simp add: bit_minus_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   190
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   191
lemma bit_minus_exp_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   192
  \<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   193
  oops
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   194
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   195
lemma bit_minus_2_iff [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   196
  \<open>bit (- 2) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n > 0\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   197
  by (simp add: bit_minus_iff bit_1_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   198
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   199
lemma not_one [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   200
  "NOT 1 = - 2"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   201
  by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   202
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   203
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   204
  apply standard
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   205
  apply (auto simp add: bit_eq_iff bit_and_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   206
  apply (simp add: bit_exp_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   207
  apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   208
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   209
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   210
sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   211
  rewrites \<open>bit.xor = (XOR)\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   212
proof -
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   213
  interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   214
    apply standard
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   215
         apply simp_all
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   216
       apply (auto simp add: bit_eq_iff bit_and_iff bit_or_iff bit_not_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   217
    apply (simp add: bit_exp_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   218
    apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   219
    done
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   220
  show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   221
    by standard
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   222
  show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close> 
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   223
    apply (auto simp add: fun_eq_iff bit.xor_def bit_eq_iff bit_and_iff bit_or_iff bit_not_iff bit_xor_iff)
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   224
         apply (simp_all add: bit_exp_iff, simp_all add: bit_def)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   225
        apply (metis local.bit_exp_iff local.bits_div_by_0)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   226
       apply (metis local.bit_exp_iff local.bits_div_by_0)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   227
    done
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   228
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   229
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   230
lemma push_bit_minus:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   231
  \<open>push_bit n (- a) = - push_bit n a\<close>
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   232
  by (simp add: push_bit_eq_mult)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   233
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   234
lemma take_bit_not_take_bit:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   235
  \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   236
  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   237
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   238
lemma take_bit_not_iff:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   239
  "take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   240
  apply (simp add: bit_eq_iff bit_not_iff bit_take_bit_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   241
  apply (simp add: bit_exp_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   242
  apply (use local.exp_eq_0_imp_not_bit in blast)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   243
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   244
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   245
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   246
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   247
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   248
subsubsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   249
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   250
locale zip_nat = single: abel_semigroup f
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   251
    for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70) +
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   252
  assumes end_of_bits: "\<not> False \<^bold>* False"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   253
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   254
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   255
lemma False_P_imp:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   256
  "False \<^bold>* True \<and> P" if "False \<^bold>* P"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   257
  using that end_of_bits by (cases P) simp_all
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   258
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   259
function F :: "nat \<Rightarrow> nat \<Rightarrow> nat"  (infixl "\<^bold>\<times>" 70)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   260
  where "m \<^bold>\<times> n = (if m = 0 \<and> n = 0 then 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   261
    else of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   262
  by auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   263
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   264
termination
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   265
  by (relation "measure (case_prod (+))") auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   266
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   267
lemma zero_left_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   268
  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   269
  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   270
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   271
lemma zero_right_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   272
  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   273
  by (induction m rule: nat_bit_induct) (simp_all add: end_of_bits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   274
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   275
lemma simps [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   276
  "0 \<^bold>\<times> 0 = 0"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   277
  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   278
  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   279
  "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   280
  by (simp_all only: zero_left_eq zero_right_eq) simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   281
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   282
lemma rec:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   283
  "m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   284
  by (cases "m = 0 \<and> n = 0") (auto simp add: end_of_bits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   285
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   286
declare F.simps [simp del]
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   287
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   288
sublocale abel_semigroup F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   289
proof
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   290
  show "m \<^bold>\<times> n \<^bold>\<times> q = m \<^bold>\<times> (n \<^bold>\<times> q)" for m n q :: nat
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   291
  proof (induction m arbitrary: n q rule: nat_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   292
    case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   293
    show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   294
      by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   295
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   296
    case (even m)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   297
    with rec [of "2 * m"] rec [of _ q] show ?case
71181
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
   298
      by (cases "even n") (auto simp add: ac_simps dest: False_P_imp)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   299
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   300
    case (odd m)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   301
    with rec [of "Suc (2 * m)"] rec [of _ q] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   302
      by (cases "even n"; cases "even q")
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   303
        (auto dest: False_P_imp simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   304
  qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   305
  show "m \<^bold>\<times> n = n \<^bold>\<times> m" for m n :: nat
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   306
  proof (induction m arbitrary: n rule: nat_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   307
    case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   308
    show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   309
      by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   310
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   311
    case (even m)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   312
    with rec [of "2 * m" n] rec [of n "2 * m"] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   313
      by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   314
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   315
    case (odd m)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   316
    with rec [of "Suc (2 * m)" n] rec [of n "Suc (2 * m)"] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   317
      by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   318
  qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   319
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   320
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   321
lemma self [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   322
  "n \<^bold>\<times> n = of_bool (True \<^bold>* True) * n"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   323
  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   324
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   325
lemma even_iff [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   326
  "even (m \<^bold>\<times> n) \<longleftrightarrow> \<not> (odd m \<^bold>* odd n)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   327
proof (induction m arbitrary: n rule: nat_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   328
  case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   329
  show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   330
    by (cases "even n") (simp_all add: end_of_bits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   331
next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   332
  case (even m)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   333
  then show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   334
    by (simp add: rec [of "2 * m"])
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   335
next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   336
  case (odd m)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   337
  then show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   338
    by (simp add: rec [of "Suc (2 * m)"])
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   339
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   340
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   341
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   342
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   343
instantiation nat :: semiring_bit_operations
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   344
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   345
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   346
global_interpretation and_nat: zip_nat "(\<and>)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   347
  defines and_nat = and_nat.F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   348
  by standard auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   349
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   350
global_interpretation and_nat: semilattice "(AND) :: nat \<Rightarrow> nat \<Rightarrow> nat"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   351
proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   352
  show "n AND n = n" for n :: nat
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   353
    by (simp add: and_nat.self)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   354
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   355
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   356
declare and_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   357
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   358
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   359
global_interpretation or_nat: zip_nat "(\<or>)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   360
  defines or_nat = or_nat.F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   361
  by standard auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   362
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   363
global_interpretation or_nat: semilattice "(OR) :: nat \<Rightarrow> nat \<Rightarrow> nat"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   364
proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   365
  show "n OR n = n" for n :: nat
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   366
    by (simp add: or_nat.self)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   367
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   368
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   369
declare or_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   370
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   371
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   372
global_interpretation xor_nat: zip_nat "(\<noteq>)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   373
  defines xor_nat = xor_nat.F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   374
  by standard auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   375
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   376
declare xor_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   377
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   378
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   379
instance proof
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   380
  fix m n q :: nat
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   381
  show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   382
  proof (rule sym, induction q arbitrary: m n)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   383
    case 0
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   384
    then show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   385
      by (simp add: and_nat.even_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   386
  next
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   387
    case (Suc q)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   388
    with and_nat.rec [of m n] show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   389
      by simp
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   390
  qed
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   391
  show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   392
  proof (rule sym, induction q arbitrary: m n)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   393
    case 0
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   394
    then show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   395
      by (simp add: or_nat.even_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   396
  next
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   397
    case (Suc q)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   398
    with or_nat.rec [of m n] show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   399
      by simp
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   400
  qed
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   401
  show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   402
  proof (rule sym, induction q arbitrary: m n)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   403
    case 0
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   404
    then show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   405
      by (simp add: xor_nat.even_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   406
  next
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   407
    case (Suc q)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   408
    with xor_nat.rec [of m n] show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   409
      by simp
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   410
  qed
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   411
qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   412
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   413
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   414
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   415
global_interpretation or_nat: semilattice_neutr "(OR)" "0 :: nat"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   416
  by standard simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   417
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   418
global_interpretation xor_nat: comm_monoid "(XOR)" "0 :: nat"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   419
  by standard simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   420
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   421
lemma Suc_0_and_eq [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   422
  "Suc 0 AND n = n mod 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   423
  by (cases n) auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   424
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   425
lemma and_Suc_0_eq [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   426
  "n AND Suc 0 = n mod 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   427
  using Suc_0_and_eq [of n] by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   428
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   429
lemma Suc_0_or_eq [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   430
  "Suc 0 OR n = n + of_bool (even n)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   431
  by (cases n) (simp_all add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   432
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   433
lemma or_Suc_0_eq [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   434
  "n OR Suc 0 = n + of_bool (even n)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   435
  using Suc_0_or_eq [of n] by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   436
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   437
lemma Suc_0_xor_eq [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   438
  "Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   439
  by (cases n) (simp_all add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   440
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   441
lemma xor_Suc_0_eq [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   442
  "n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   443
  using Suc_0_xor_eq [of n] by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   444
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   445
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   446
subsubsection \<open>Instance \<^typ>\<open>int\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   447
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   448
abbreviation (input) complement :: "int \<Rightarrow> int"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   449
  where "complement k \<equiv> - k - 1"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   450
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   451
lemma complement_half:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   452
  "complement (k * 2) div 2 = complement k"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   453
  by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   454
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   455
lemma complement_div_2:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   456
  "complement (k div 2) = complement k div 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   457
  by linarith
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   458
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   459
locale zip_int = single: abel_semigroup f
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   460
  for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   461
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   462
 
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   463
lemma False_False_imp_True_True:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   464
  "True \<^bold>* True" if "False \<^bold>* False"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   465
proof (rule ccontr)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   466
  assume "\<not> True \<^bold>* True"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   467
  with that show False
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   468
    using single.assoc [of False True True]
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   469
    by (cases "False \<^bold>* True") simp_all
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   470
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   471
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   472
function F :: "int \<Rightarrow> int \<Rightarrow> int"  (infixl "\<^bold>\<times>" 70)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   473
  where "k \<^bold>\<times> l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   474
    then - of_bool (odd k \<^bold>* odd l)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   475
    else of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   476
  by auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   477
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   478
termination
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   479
  by (relation "measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))") auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   480
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   481
lemma zero_left_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   482
  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   483
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   484
     | (False, True) \<Rightarrow> l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   485
     | (True, False) \<Rightarrow> complement l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   486
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   487
  by (induction l rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   488
   (simp_all split: bool.split) 
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   489
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   490
lemma minus_left_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   491
  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   492
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   493
     | (False, True) \<Rightarrow> l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   494
     | (True, False) \<Rightarrow> complement l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   495
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   496
  by (induction l rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   497
   (simp_all split: bool.split) 
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   498
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   499
lemma zero_right_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   500
  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   501
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   502
     | (False, True) \<Rightarrow> k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   503
     | (True, False) \<Rightarrow> complement k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   504
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   505
  by (induction k rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   506
    (simp_all add: ac_simps split: bool.split)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   507
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   508
lemma minus_right_eq:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   509
  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   510
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   511
     | (False, True) \<Rightarrow> k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   512
     | (True, False) \<Rightarrow> complement k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   513
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   514
  by (induction k rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   515
    (simp_all add: ac_simps split: bool.split)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   516
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   517
lemma simps [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   518
  "0 \<^bold>\<times> 0 = - of_bool (False \<^bold>* False)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   519
  "- 1 \<^bold>\<times> 0 = - of_bool (True \<^bold>* False)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   520
  "0 \<^bold>\<times> - 1 = - of_bool (False \<^bold>* True)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   521
  "- 1 \<^bold>\<times> - 1 = - of_bool (True \<^bold>* True)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   522
  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   523
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   524
     | (False, True) \<Rightarrow> l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   525
     | (True, False) \<Rightarrow> complement l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   526
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   527
  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   528
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   529
     | (False, True) \<Rightarrow> l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   530
     | (True, False) \<Rightarrow> complement l
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   531
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   532
  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   533
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   534
     | (False, True) \<Rightarrow> k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   535
     | (True, False) \<Rightarrow> complement k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   536
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   537
  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   538
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   539
     | (False, True) \<Rightarrow> k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   540
     | (True, False) \<Rightarrow> complement k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   541
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   542
  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> l \<noteq> - 1
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   543
    \<Longrightarrow> k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   544
  by simp_all[4] (simp_all only: zero_left_eq minus_left_eq zero_right_eq minus_right_eq, simp)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   545
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   546
declare F.simps [simp del]
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   547
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   548
lemma rec:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   549
  "k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   550
  by (cases "k \<in> {0, - 1} \<and> l \<in> {0, - 1}") (auto simp add: ac_simps F.simps [of k l] split: bool.split)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   551
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   552
sublocale abel_semigroup F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   553
proof
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   554
  show "k \<^bold>\<times> l \<^bold>\<times> r = k \<^bold>\<times> (l \<^bold>\<times> r)" for k l r :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   555
  proof (induction k arbitrary: l r rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   556
    case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   557
    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "False \<^bold>* False" "\<not> False \<^bold>* True"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   558
    proof (induction l arbitrary: r rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   559
      case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   560
      from that show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   561
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   562
    next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   563
      case minus
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   564
      from that show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   565
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   566
    next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   567
      case (even l)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   568
      with that rec [of _ r] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   569
        by (cases "even r")
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   570
          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   571
    next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   572
      case (odd l)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   573
      moreover have "- l - 1 = - 1 - l"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   574
        by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   575
      ultimately show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   576
        using that rec [of _ r]
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   577
        by (cases "even r")
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   578
          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   579
    qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   580
    then show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   581
      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   582
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   583
    case minus
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   584
    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "\<not> True \<^bold>* True" "False \<^bold>* True"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   585
    proof (induction l arbitrary: r rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   586
      case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   587
      from that show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   588
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   589
    next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   590
      case minus
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   591
      from that show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   592
        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   593
    next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   594
      case (even l)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   595
      with that rec [of _ r] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   596
        by (cases "even r")
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   597
          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   598
    next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   599
      case (odd l)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   600
      moreover have "- l - 1 = - 1 - l"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   601
        by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   602
      ultimately show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   603
        using that rec [of _ r]
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   604
        by (cases "even r")
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   605
          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   606
    qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   607
    then show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   608
      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   609
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   610
    case (even k)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   611
    with rec [of "k * 2"] rec [of _ r] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   612
      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   613
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   614
    case (odd k)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   615
    with rec [of "1 + k * 2"] rec [of _ r] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   616
      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   617
  qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   618
  show "k \<^bold>\<times> l = l \<^bold>\<times> k" for k l :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   619
  proof (induction k arbitrary: l rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   620
    case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   621
    show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   622
      by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   623
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   624
    case minus
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   625
    show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   626
      by simp
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   627
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   628
    case (even k)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   629
    with rec [of "k * 2" l] rec [of l "k * 2"] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   630
      by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   631
  next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   632
    case (odd k)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   633
    with rec [of "k * 2 + 1" l] rec [of l "k * 2 + 1"] show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   634
      by (simp add: ac_simps)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   635
  qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   636
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   637
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   638
lemma self [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   639
  "k \<^bold>\<times> k = (case (False \<^bold>* False, True \<^bold>* True)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   640
    of (False, False) \<Rightarrow> 0
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   641
     | (False, True) \<Rightarrow> k
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   642
     | (True, True) \<Rightarrow> - 1)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   643
  by (induction k rule: int_bit_induct) (auto simp add: False_False_imp_True_True split: bool.split)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   644
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   645
lemma even_iff [simp]:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   646
  "even (k \<^bold>\<times> l) \<longleftrightarrow> \<not> (odd k \<^bold>* odd l)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   647
proof (induction k arbitrary: l rule: int_bit_induct)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   648
  case zero
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   649
  show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   650
    by (cases "even l") (simp_all split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   651
next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   652
  case minus
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   653
  show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   654
    by (cases "even l") (simp_all split: bool.splits)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   655
next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   656
  case (even k)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   657
  then show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   658
    by (simp add: rec [of "k * 2"])
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   659
next
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   660
  case (odd k)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   661
  then show ?case
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   662
    by (simp add: rec [of "1 + k * 2"])
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   663
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   664
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   665
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   666
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   667
instantiation int :: ring_bit_operations
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   668
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   669
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   670
definition not_int :: "int \<Rightarrow> int"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   671
  where "not_int = complement"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   672
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   673
global_interpretation and_int: zip_int "(\<and>)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   674
  defines and_int = and_int.F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   675
  by standard
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   676
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   677
declare and_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   678
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   679
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   680
global_interpretation and_int: semilattice "(AND) :: int \<Rightarrow> int \<Rightarrow> int"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   681
proof (rule semilattice.intro, fact and_int.abel_semigroup_axioms, standard)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   682
  show "k AND k = k" for k :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   683
    by (simp add: and_int.self)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   684
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   685
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   686
global_interpretation or_int: zip_int "(\<or>)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   687
  defines or_int = or_int.F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   688
  by standard
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   689
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   690
declare or_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   691
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   692
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   693
global_interpretation or_int: semilattice "(OR) :: int \<Rightarrow> int \<Rightarrow> int"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   694
proof (rule semilattice.intro, fact or_int.abel_semigroup_axioms, standard)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   695
  show "k OR k = k" for k :: int
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   696
    by (simp add: or_int.self)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   697
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   698
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   699
global_interpretation xor_int: zip_int "(\<noteq>)"
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   700
  defines xor_int = xor_int.F
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   701
  by standard
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   702
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   703
declare xor_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   704
  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   705
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   706
lemma bit_not_iff_int:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   707
  \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   708
    for k :: int
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   709
    by (induction n arbitrary: k)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   710
      (simp_all add: not_int_def flip: complement_div_2)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   711
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   712
instance proof
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   713
  fix k l :: int and n :: nat
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   714
  show \<open>- k = NOT (k - 1)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   715
    by (simp add: not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   716
  show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   717
  proof (rule sym, induction n arbitrary: k l)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   718
    case 0
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   719
    then show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   720
      by (simp add: and_int.even_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   721
  next
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   722
    case (Suc n)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   723
    with and_int.rec [of k l] show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   724
      by simp
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   725
  qed
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   726
  show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   727
  proof (rule sym, induction n arbitrary: k l)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   728
    case 0
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   729
    then show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   730
      by (simp add: or_int.even_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   731
  next
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   732
    case (Suc n)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   733
    with or_int.rec [of k l] show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   734
      by simp
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   735
  qed
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   736
  show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   737
  proof (rule sym, induction n arbitrary: k l)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   738
    case 0
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   739
    then show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   740
      by (simp add: xor_int.even_iff)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   741
  next
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   742
    case (Suc n)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   743
    with xor_int.rec [of k l] show ?case
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   744
      by simp
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   745
  qed
71195
d50a718ccf35 tuned material
haftmann
parents: 71186
diff changeset
   746
qed (simp_all add: minus_1_div_exp_eq_int bit_not_iff_int)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   747
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   748
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   749
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   750
lemma not_int_div_2:
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   751
  "NOT k div 2 = NOT (k div 2)" for k :: int
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   752
  by (simp add: complement_div_2 not_int_def)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   753
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   754
lemma not_int_rec [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   755
  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   756
  by (auto simp add: not_int_def elim: oddE)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   757
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   758
end