src/HOL/ex/Bit_Operations.thy
author haftmann
Sun, 09 Feb 2020 10:46:32 +0000
changeset 71424 e83fe2c31088
parent 71420 572ab9e64e18
child 71426 745e518d3d0b
permissions -rw-r--r--
rule concerning bit (push_bit ...)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     1
(*  Author:  Florian Haftmann, TUM
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     2
*)
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
     3
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
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    12
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
    13
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    14
class semiring_bit_operations = semiring_bit_shifts +
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    15
  fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr "AND" 64)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    16
    and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr "OR"  59)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    17
    and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>  (infixr "XOR" 59)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    18
  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
    19
    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
    20
    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
    21
begin
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
text \<open>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    24
  We want the bitwise operations to bind slightly weaker
71094
a197532693a5 bit shifts as class operations
haftmann
parents: 71042
diff changeset
    25
  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
    26
  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
    27
  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
    28
  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
    29
\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    30
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    31
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
    32
  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
    33
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    34
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
    35
  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
    36
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    37
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
    38
  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
    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 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
    41
  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
    42
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    43
text \<open>
71181
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    44
  Having 
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
    45
  \<^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
    46
  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
    47
\<close>
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
    48
71181
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    49
lemma stable_imp_drop_eq:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    50
  \<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
    51
  by (induction n) (simp_all add: that)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    52
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    53
lemma map_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    54
  \<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
    55
  by (simp add: map_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    56
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    57
lemma map_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    58
  \<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
    59
  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
    60
    elim: evenE oddE)
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 set_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    63
  \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    64
  by (simp add: set_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 set_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    67
  \<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
    68
  by (simp add: set_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    69
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    70
lemma unset_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    71
  \<open>unset_bit 0 a = 2 * (a div 2)\<close>
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    72
  by (simp add: unset_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    73
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    74
lemma unset_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    75
  \<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
    76
  by (simp add: unset_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    77
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    78
lemma flip_bit_0 [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    79
  \<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
    80
  by (simp add: flip_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    81
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    82
lemma flip_bit_Suc [simp]:
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    83
  \<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
    84
  by (simp add: flip_bit_def)
8331063570d6 bit accessor and fundamental properties
haftmann
parents: 71095
diff changeset
    85
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    86
sublocale "and": semilattice \<open>(AND)\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    87
  by standard (auto simp add: bit_eq_iff bit_and_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    88
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    89
sublocale or: semilattice_neutr \<open>(OR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    90
  by standard (auto simp add: bit_eq_iff bit_or_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    91
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    92
sublocale xor: comm_monoid \<open>(XOR)\<close> 0
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    93
  by standard (auto simp add: bit_eq_iff bit_xor_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
    94
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    95
lemma zero_and_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    96
  "0 AND a = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    97
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    98
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
    99
lemma and_zero_eq [simp]:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   100
  "a AND 0 = 0"
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   101
  by (simp add: bit_eq_iff bit_and_iff)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   102
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   103
lemma one_and_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   104
  "1 AND a = of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   105
  by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   106
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   107
lemma and_one_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   108
  "a AND 1 = of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   109
  using one_and_eq [of a] by (simp add: ac_simps)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   110
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   111
lemma one_or_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   112
  "1 OR a = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   113
  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
   114
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   115
lemma or_one_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   116
  "a OR 1 = a + of_bool (even a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   117
  using one_or_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   118
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   119
lemma one_xor_eq [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   120
  "1 XOR a = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   121
  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
   122
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   123
lemma xor_one_eq [simp]:
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   124
  "a XOR 1 = a + of_bool (even a) - of_bool (odd a)"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   125
  using one_xor_eq [of a] by (simp add: ac_simps)
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   126
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   127
lemma take_bit_and [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   128
  \<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
   129
  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
   130
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   131
lemma take_bit_or [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   132
  \<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
   133
  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
   134
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   135
lemma take_bit_xor [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   136
  \<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
   137
  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
   138
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   139
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   140
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   141
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
   142
  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   143
  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
   144
  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
   145
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   146
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   147
text \<open>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   148
  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
   149
  definitional class operation.  Note that \<^const>\<open>not\<close> has no
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   150
  sensible definition for unlimited but only positive bit strings
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   151
  (type \<^typ>\<open>nat\<close>).
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   152
\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   153
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   154
lemma bits_minus_1_mod_2_eq [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   155
  \<open>(- 1) mod 2 = 1\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   156
  by (simp add: mod_2_eq_odd)
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   157
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   158
lemma not_eq_complement:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   159
  \<open>NOT a = - a - 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   160
  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
   161
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   162
lemma minus_eq_not_plus_1:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   163
  \<open>- a = NOT a + 1\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   164
  using not_eq_complement [of a] by simp
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   165
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   166
lemma bit_minus_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   167
  \<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
   168
  by (simp add: minus_eq_not_minus_1 bit_not_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   169
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   170
lemma even_not_iff [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   171
  "even (NOT a) \<longleftrightarrow> odd a"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   172
  using bit_not_iff [of a 0] by auto
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   173
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   174
lemma bit_not_exp_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   175
  \<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
   176
  by (auto simp add: bit_not_iff bit_exp_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   177
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   178
lemma bit_minus_1_iff [simp]:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   179
  \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   180
  by (simp add: bit_minus_iff)
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   181
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   182
lemma bit_minus_exp_iff:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   183
  \<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
   184
  oops
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   185
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   186
lemma bit_minus_2_iff [simp]:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   187
  \<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
   188
  by (simp add: bit_minus_iff bit_1_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   189
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   190
lemma not_one [simp]:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   191
  "NOT 1 = - 2"
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   192
  by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   193
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   194
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   195
  apply standard
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   196
  apply (auto simp add: bit_eq_iff bit_and_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   197
  apply (simp add: bit_exp_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   198
  apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   199
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   200
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   201
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
   202
  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
   203
proof -
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   204
  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
   205
    apply standard
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   206
         apply simp_all
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   207
       apply (auto simp add: bit_eq_iff bit_and_iff bit_or_iff bit_not_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   208
    apply (simp add: bit_exp_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   209
    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
   210
    done
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   211
  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
   212
    by standard
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   213
  show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close> 
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   214
    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
   215
         apply (simp_all add: bit_exp_iff, simp_all add: bit_def)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   216
        apply (metis local.bit_exp_iff local.bits_div_by_0)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   217
       apply (metis local.bit_exp_iff local.bits_div_by_0)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   218
    done
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   219
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   220
71412
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   221
lemma push_bit_minus:
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   222
  \<open>push_bit n (- a) = - push_bit n a\<close>
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   223
  by (simp add: push_bit_eq_mult)
96d126844adc more theorems
haftmann
parents: 71409
diff changeset
   224
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   225
lemma take_bit_not_take_bit:
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   226
  \<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
   227
  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
   228
71418
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   229
lemma take_bit_not_iff:
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   230
  "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
   231
  apply (simp add: bit_eq_iff bit_not_iff bit_take_bit_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   232
  apply (simp add: bit_exp_iff)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   233
  apply (use local.exp_eq_0_imp_not_bit in blast)
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   234
  done
bd9d27ccb3a3 more theorems
haftmann
parents: 71413
diff changeset
   235
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   236
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   237
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   238
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   239
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
   240
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   241
locale zip_nat = single: abel_semigroup f
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   242
    for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl \<open>\<^bold>*\<close> 70) +
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   243
  assumes end_of_bits: \<open>\<not> False \<^bold>* False\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   244
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   245
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   246
function F :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>  (infixl \<open>\<^bold>\<times>\<close> 70)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   247
  where \<open>m \<^bold>\<times> n = (if m = 0 \<and> n = 0 then 0
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   248
    else of_bool (odd m \<^bold>* odd n) + 2 * ((m div 2) \<^bold>\<times> (n div 2)))\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   249
  by auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   250
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   251
termination
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   252
  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
   253
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   254
declare F.simps [simp del]
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   255
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   256
lemma rec:
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   257
  "m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   258
proof (cases \<open>m = 0 \<and> n = 0\<close>)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   259
  case True
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   260
  then have \<open>m \<^bold>\<times> n = 0\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   261
    using True by (simp add: F.simps [of 0 0])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   262
  moreover have \<open>(m div 2) \<^bold>\<times> (n div 2) = m \<^bold>\<times> n\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   263
    using True by simp
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   264
  ultimately show ?thesis
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   265
    using True by (simp add: end_of_bits)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   266
next
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   267
  case False
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   268
  then show ?thesis
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   269
    by (auto simp add: ac_simps F.simps [of m n])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   270
qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   271
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   272
lemma bit_eq_iff:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   273
  \<open>bit (m \<^bold>\<times> n) q \<longleftrightarrow> bit m q \<^bold>* bit n q\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   274
proof (induction q arbitrary: m n)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   275
  case 0
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   276
  then show ?case
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   277
    by (simp add: rec [of m n])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   278
next
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   279
  case (Suc n)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   280
  then show ?case
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   281
    by (simp add: rec [of m n])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   282
qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   283
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   284
sublocale abel_semigroup F
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   285
  by standard (simp_all add: Parity.bit_eq_iff bit_eq_iff ac_simps)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   286
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   287
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   288
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   289
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
   290
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   291
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   292
global_interpretation and_nat: zip_nat \<open>(\<and>)\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   293
  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
   294
  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
   295
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   296
global_interpretation and_nat: semilattice \<open>(AND) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   297
proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard)
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   298
  show \<open>n AND n = n\<close> for n :: nat
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   299
    by (simp add: bit_eq_iff and_nat.bit_eq_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   300
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   301
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   302
global_interpretation or_nat: zip_nat \<open>(\<or>)\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   303
  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
   304
  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
   305
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   306
global_interpretation or_nat: semilattice \<open>(OR) :: nat \<Rightarrow> nat \<Rightarrow> nat\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   307
proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard)
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   308
  show \<open>n OR n = n\<close> for n :: nat
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   309
    by (simp add: bit_eq_iff or_nat.bit_eq_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   310
qed
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   311
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   312
global_interpretation xor_nat: zip_nat \<open>(\<noteq>)\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   313
  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
   314
  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
   315
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   316
instance proof
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   317
  fix m n q :: nat
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   318
  show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   319
    by (fact and_nat.bit_eq_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   320
  show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   321
    by (fact or_nat.bit_eq_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   322
  show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   323
    by (fact xor_nat.bit_eq_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   324
qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   325
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   326
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   327
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   328
lemma Suc_0_and_eq [simp]:
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   329
  \<open>Suc 0 AND n = of_bool (odd n)\<close>
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   330
  using one_and_eq [of n] 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
   331
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   332
lemma and_Suc_0_eq [simp]:
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   333
  \<open>n AND Suc 0 = of_bool (odd n)\<close>
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   334
  using and_one_eq [of n] 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
   335
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   336
lemma Suc_0_or_eq [simp]:
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   337
  \<open>Suc 0 OR n = n + of_bool (even n)\<close>
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   338
  using one_or_eq [of n] 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
   339
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   340
lemma or_Suc_0_eq [simp]:
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   341
  \<open>n OR Suc 0 = n + of_bool (even n)\<close>
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   342
  using or_one_eq [of n] 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
   343
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   344
lemma Suc_0_xor_eq [simp]:
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   345
  \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close>
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   346
  using one_xor_eq [of n] 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
   347
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   348
lemma xor_Suc_0_eq [simp]:
71419
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   349
  \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close>
1d8e914e04d6 simplified logical constructions
haftmann
parents: 71418
diff changeset
   350
  using xor_one_eq [of n] 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
   351
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   352
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   353
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
   354
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   355
locale zip_int = single: abel_semigroup f
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   356
  for f :: \<open>bool \<Rightarrow> bool \<Rightarrow> bool\<close>  (infixl \<open>\<^bold>*\<close> 70)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   357
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   358
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   359
function F :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>  (infixl \<open>\<^bold>\<times>\<close> 70)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   360
  where \<open>k \<^bold>\<times> l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   361
    then - of_bool (odd k \<^bold>* odd l)
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   362
    else of_bool (odd k \<^bold>* odd l) + 2 * ((k div 2) \<^bold>\<times> (l div 2)))\<close>
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   363
  by auto
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   364
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   365
termination
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   366
  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
   367
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   368
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
   369
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   370
lemma rec:
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   371
  \<open>k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + 2 * ((k div 2) \<^bold>\<times> (l div 2))\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   372
proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   373
  case True
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   374
  then have \<open>(k div 2) \<^bold>\<times> (l div 2) = k \<^bold>\<times> l\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   375
    by auto
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   376
  moreover have \<open>of_bool (odd k \<^bold>* odd l) = - (k \<^bold>\<times> l)\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   377
    using True by (simp add: F.simps [of k l])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   378
  ultimately show ?thesis by simp
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   379
next
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   380
  case False
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   381
  then show ?thesis
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   382
    by (auto simp add: ac_simps F.simps [of k l])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   383
qed
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   384
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   385
lemma bit_eq_iff:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   386
  \<open>bit (k \<^bold>\<times> l) n \<longleftrightarrow> bit k n \<^bold>* bit l n\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   387
proof (induction n arbitrary: k l)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   388
  case 0
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   389
  then show ?case
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   390
    by (simp add: rec [of k l])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   391
next
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   392
  case (Suc n)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   393
  then show ?case
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   394
    by (simp add: rec [of k l])
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   395
qed
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   396
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   397
sublocale abel_semigroup F
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   398
  by standard (simp_all add: Parity.bit_eq_iff bit_eq_iff ac_simps)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   399
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   400
end
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   401
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   402
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
   403
begin
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   404
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   405
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
   406
  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
   407
  by standard
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   408
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   409
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
   410
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
   411
  show "k AND k = k" for k :: int
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   412
    by (simp add: bit_eq_iff and_int.bit_eq_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   413
qed
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_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
   416
  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
   417
  by standard
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   418
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   419
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
   420
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
   421
  show "k OR k = k" for k :: int
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   422
    by (simp add: bit_eq_iff or_int.bit_eq_iff)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   423
qed
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
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
   426
  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
   427
  by standard
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   428
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   429
definition not_int :: \<open>int \<Rightarrow> int\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   430
  where \<open>not_int k = - k - 1\<close>
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   431
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   432
lemma not_int_rec:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   433
  "NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   434
  by (auto simp add: not_int_def elim: oddE)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   435
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   436
lemma even_not_iff_int:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   437
  \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   438
  by (simp add: not_int_def)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   439
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   440
lemma not_int_div_2:
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   441
  \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   442
  by (simp add: not_int_def)
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   443
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   444
lemma bit_not_iff_int:
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   445
  \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   446
    for k :: int
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   447
  by (induction n arbitrary: k) (simp_all add: not_int_div_2 even_not_iff_int)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   448
71042
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   449
instance proof
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   450
  fix k l :: int and n :: nat
71409
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   451
  show \<open>- k = NOT (k - 1)\<close>
0bb0cb558bf9 sketches of ideas still to come
haftmann
parents: 71195
diff changeset
   452
    by (simp add: not_int_def)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   453
  show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   454
    by (fact and_int.bit_eq_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   455
  show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   456
    by (fact or_int.bit_eq_iff)
71186
3d35e12999ba characterization of typical bit operations
haftmann
parents: 71181
diff changeset
   457
  show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
71420
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   458
    by (fact xor_int.bit_eq_iff)
572ab9e64e18 simplified logical constructions
haftmann
parents: 71419
diff changeset
   459
qed (simp_all add: 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
   460
400e9512f1d3 proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff changeset
   461
end
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
end