| author | wenzelm | 
| Thu, 06 Jun 2024 22:03:20 +0200 | |
| changeset 80271 | 198fc882ec0f | 
| parent 79072 | a91050cd5c93 | 
| permissions | -rw-r--r-- | 
| 70342 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
69593diff
changeset | 1 | (* Title: HOL/Library/Z2.thy | 
| 41959 | 2 | Author: Brian Huffman | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 3 | *) | 
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 4 | |
| 60500 | 5 | section \<open>The Field of Integers mod 2\<close> | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 6 | |
| 70342 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
69593diff
changeset | 7 | theory Z2 | 
| 74101 | 8 | imports Main | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 9 | begin | 
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 10 | |
| 70342 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
69593diff
changeset | 11 | text \<open> | 
| 74101 | 12 | Note that in most cases \<^typ>\<open>bool\<close> is appropriate when a binary type is needed; the | 
| 70351 | 13 | type provided here, for historical reasons named \<^text>\<open>bit\<close>, is only needed if proper | 
| 70342 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
69593diff
changeset | 14 | field operations are required. | 
| 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
69593diff
changeset | 15 | \<close> | 
| 
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
 haftmann parents: 
69593diff
changeset | 16 | |
| 71988 | 17 | typedef bit = \<open>UNIV :: bool set\<close> .. | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 18 | |
| 71988 | 19 | instantiation bit :: zero_neq_one | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 20 | begin | 
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 21 | |
| 71988 | 22 | definition zero_bit :: bit | 
| 23 | where \<open>0 = Abs_bit False\<close> | |
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 24 | |
| 71988 | 25 | definition one_bit :: bit | 
| 26 | where \<open>1 = Abs_bit True\<close> | |
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 27 | |
| 71988 | 28 | instance | 
| 29 | by standard (simp add: zero_bit_def one_bit_def Abs_bit_inject) | |
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 30 | |
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 31 | end | 
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 32 | |
| 71988 | 33 | free_constructors case_bit for \<open>0::bit\<close> | \<open>1::bit\<close> | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 34 | proof - | 
| 71988 | 35 | fix P :: bool | 
| 36 | fix a :: bit | |
| 37 | assume \<open>a = 0 \<Longrightarrow> P\<close> and \<open>a = 1 \<Longrightarrow> P\<close> | |
| 38 | then show P | |
| 39 | by (cases a) (auto simp add: zero_bit_def one_bit_def Abs_bit_inject) | |
| 40 | qed simp | |
| 41 | ||
| 42 | lemma bit_not_zero_iff [simp]: | |
| 43 | \<open>a \<noteq> 0 \<longleftrightarrow> a = 1\<close> for a :: bit | |
| 44 | by (cases a) simp_all | |
| 45 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 46 | lemma bit_not_one_iff [simp]: | 
| 71988 | 47 | \<open>a \<noteq> 1 \<longleftrightarrow> a = 0\<close> for a :: bit | 
| 48 | by (cases a) simp_all | |
| 49 | ||
| 50 | instantiation bit :: semidom_modulo | |
| 51 | begin | |
| 52 | ||
| 53 | definition plus_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 54 | where \<open>a + b = Abs_bit (Rep_bit a \<noteq> Rep_bit b)\<close> | |
| 55 | ||
| 56 | definition minus_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 57 | where [simp]: \<open>minus_bit = plus\<close> | |
| 58 | ||
| 59 | definition times_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 60 | where \<open>a * b = Abs_bit (Rep_bit a \<and> Rep_bit b)\<close> | |
| 61 | ||
| 62 | definition divide_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 63 | where [simp]: \<open>divide_bit = times\<close> | |
| 64 | ||
| 65 | definition modulo_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 66 | where \<open>a mod b = Abs_bit (Rep_bit a \<and> \<not> Rep_bit b)\<close> | |
| 67 | ||
| 68 | instance | |
| 69 | by standard | |
| 70 | (auto simp flip: Rep_bit_inject | |
| 71 | simp add: zero_bit_def one_bit_def plus_bit_def times_bit_def modulo_bit_def Abs_bit_inverse Rep_bit_inverse) | |
| 72 | ||
| 73 | end | |
| 74 | ||
| 75 | lemma bit_2_eq_0 [simp]: | |
| 76 | \<open>2 = (0::bit)\<close> | |
| 77 | by (simp flip: one_add_one add: zero_bit_def plus_bit_def) | |
| 78 | ||
| 79 | instance bit :: semiring_parity | |
| 80 | apply standard | |
| 81 | apply (auto simp flip: Rep_bit_inject simp add: modulo_bit_def Abs_bit_inverse Rep_bit_inverse) | |
| 82 | apply (auto simp add: zero_bit_def one_bit_def Abs_bit_inverse Rep_bit_inverse) | |
| 83 | done | |
| 84 | ||
| 85 | lemma Abs_bit_eq_of_bool [code_abbrev]: | |
| 86 | \<open>Abs_bit = of_bool\<close> | |
| 87 | by (simp add: fun_eq_iff zero_bit_def one_bit_def) | |
| 88 | ||
| 89 | lemma Rep_bit_eq_odd: | |
| 90 | \<open>Rep_bit = odd\<close> | |
| 91 | proof - | |
| 92 | have \<open>\<not> Rep_bit 0\<close> | |
| 93 | by (simp only: zero_bit_def) (subst Abs_bit_inverse, auto) | |
| 94 | then show ?thesis | |
| 95 | by (auto simp flip: Rep_bit_inject simp add: fun_eq_iff) | |
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 96 | qed | 
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 97 | |
| 71988 | 98 | lemma Rep_bit_iff_odd [code_abbrev]: | 
| 99 | \<open>Rep_bit b \<longleftrightarrow> odd b\<close> | |
| 100 | by (simp add: Rep_bit_eq_odd) | |
| 101 | ||
| 102 | lemma Not_Rep_bit_iff_even [code_abbrev]: | |
| 103 | \<open>\<not> Rep_bit b \<longleftrightarrow> even b\<close> | |
| 104 | by (simp add: Rep_bit_eq_odd) | |
| 105 | ||
| 106 | lemma Not_Not_Rep_bit [code_unfold]: | |
| 107 | \<open>\<not> \<not> Rep_bit b \<longleftrightarrow> Rep_bit b\<close> | |
| 108 | by simp | |
| 109 | ||
| 110 | code_datatype \<open>0::bit\<close> \<open>1::bit\<close> | |
| 63462 | 111 | |
| 71988 | 112 | lemma Abs_bit_code [code]: | 
| 113 | \<open>Abs_bit False = 0\<close> | |
| 114 | \<open>Abs_bit True = 1\<close> | |
| 115 | by (simp_all add: Abs_bit_eq_of_bool) | |
| 116 | ||
| 117 | lemma Rep_bit_code [code]: | |
| 118 | \<open>Rep_bit 0 \<longleftrightarrow> False\<close> | |
| 119 | \<open>Rep_bit 1 \<longleftrightarrow> True\<close> | |
| 120 | by (simp_all add: Rep_bit_eq_odd) | |
| 121 | ||
| 122 | context zero_neq_one | |
| 123 | begin | |
| 124 | ||
| 125 | abbreviation of_bit :: \<open>bit \<Rightarrow> 'a\<close> | |
| 126 | where \<open>of_bit b \<equiv> of_bool (odd b)\<close> | |
| 127 | ||
| 128 | end | |
| 53063 
8f7ac535892f
explicit conversion from and to bool, and into algebraic structures with 0 and 1
 haftmann parents: 
49834diff
changeset | 129 | |
| 71953 | 130 | context | 
| 131 | begin | |
| 132 | ||
| 71988 | 133 | qualified lemma bit_eq_iff: | 
| 134 | \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close> for a b :: bit | |
| 135 | by (cases a; cases b) simp_all | |
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 136 | |
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 137 | end | 
| 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: diff
changeset | 138 | |
| 71988 | 139 | lemma modulo_bit_unfold [simp, code]: | 
| 140 | \<open>a mod b = of_bool (odd a \<and> even b)\<close> for a b :: bit | |
| 141 | by (simp add: modulo_bit_def Abs_bit_eq_of_bool Rep_bit_eq_odd) | |
| 53063 
8f7ac535892f
explicit conversion from and to bool, and into algebraic structures with 0 and 1
 haftmann parents: 
49834diff
changeset | 142 | |
| 75962 | 143 | lemma power_bit_unfold [simp]: | 
| 71988 | 144 | \<open>a ^ n = of_bool (odd a \<or> n = 0)\<close> for a :: bit | 
| 145 | by (cases a) simp_all | |
| 71953 | 146 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 147 | instantiation bit :: field | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 148 | begin | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 149 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 150 | definition uminus_bit :: \<open>bit \<Rightarrow> bit\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 151 | where [simp]: \<open>uminus_bit = id\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 152 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 153 | definition inverse_bit :: \<open>bit \<Rightarrow> bit\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 154 | where [simp]: \<open>inverse_bit = id\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 155 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 156 | instance | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 157 | apply standard | 
| 79072 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 haftmann parents: 
75962diff
changeset | 158 | apply simp_all | 
| 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 haftmann parents: 
75962diff
changeset | 159 | apply (simp only: Z2.bit_eq_iff even_add even_zero refl) | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 160 | done | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 161 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 162 | end | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 163 | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 164 | instantiation bit :: semiring_bits | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 165 | begin | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 166 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 167 | definition bit_bit :: \<open>bit \<Rightarrow> nat \<Rightarrow> bool\<close> | 
| 71988 | 168 | where [simp]: \<open>bit_bit b n \<longleftrightarrow> odd b \<and> n = 0\<close> | 
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 169 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 170 | instance | 
| 79072 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 haftmann parents: 
75962diff
changeset | 171 | by standard | 
| 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 haftmann parents: 
75962diff
changeset | 172 | (auto intro: Abs_bit_induct simp add: Abs_bit_eq_of_bool) | 
| 71953 | 173 | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 174 | end | 
| 71953 | 175 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 176 | instantiation bit :: ring_bit_operations | 
| 71988 | 177 | begin | 
| 178 | ||
| 74097 | 179 | context | 
| 180 | includes bit_operations_syntax | |
| 181 | begin | |
| 182 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 183 | definition not_bit :: \<open>bit \<Rightarrow> bit\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 184 | where [simp]: \<open>NOT b = of_bool (even b)\<close> for b :: bit | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 185 | |
| 71988 | 186 | definition and_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 187 | where [simp]: \<open>b AND c = of_bool (odd b \<and> odd c)\<close> for b c :: bit | |
| 188 | ||
| 189 | definition or_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 190 | where [simp]: \<open>b OR c = of_bool (odd b \<or> odd c)\<close> for b c :: bit | |
| 191 | ||
| 192 | definition xor_bit :: \<open>bit \<Rightarrow> bit \<Rightarrow> bit\<close> | |
| 193 | where [simp]: \<open>b XOR c = of_bool (odd b \<noteq> odd c)\<close> for b c :: bit | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71957diff
changeset | 194 | |
| 72082 | 195 | definition mask_bit :: \<open>nat \<Rightarrow> bit\<close> | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 196 | where [simp]: \<open>mask n = (of_bool (n > 0) :: bit)\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 197 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 198 | definition set_bit_bit :: \<open>nat \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 199 | where [simp]: \<open>set_bit n b = of_bool (n = 0 \<or> odd b)\<close> for b :: bit | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 200 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 201 | definition unset_bit_bit :: \<open>nat \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 202 | where [simp]: \<open>unset_bit n b = of_bool (n > 0 \<and> odd b)\<close> for b :: bit | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 203 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 204 | definition flip_bit_bit :: \<open>nat \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 205 | where [simp]: \<open>flip_bit n b = of_bool ((n = 0) \<noteq> odd b)\<close> for b :: bit | 
| 72082 | 206 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 207 | definition push_bit_bit :: \<open>nat \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 208 | where [simp]: \<open>push_bit n b = of_bool (odd b \<and> n = 0)\<close> for b :: bit | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 209 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 210 | definition drop_bit_bit :: \<open>nat \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 211 | where [simp]: \<open>drop_bit n b = of_bool (odd b \<and> n = 0)\<close> for b :: bit | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 212 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 213 | definition take_bit_bit :: \<open>nat \<Rightarrow> bit \<Rightarrow> bit\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 214 | where [simp]: \<open>take_bit n b = of_bool (odd b \<and> n > 0)\<close> for b :: bit | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 215 | |
| 74097 | 216 | end | 
| 217 | ||
| 71953 | 218 | instance | 
| 79072 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 haftmann parents: 
75962diff
changeset | 219 | by standard auto | 
| 71988 | 220 | |
| 221 | end | |
| 222 | ||
| 223 | lemma add_bit_eq_xor [simp, code]: | |
| 74097 | 224 | \<open>(+) = (Bit_Operations.xor :: bit \<Rightarrow> _)\<close> | 
| 71988 | 225 | by (auto simp add: fun_eq_iff) | 
| 226 | ||
| 227 | lemma mult_bit_eq_and [simp, code]: | |
| 74097 | 228 | \<open>(*) = (Bit_Operations.and :: bit \<Rightarrow> _)\<close> | 
| 71988 | 229 | by (simp add: fun_eq_iff) | 
| 230 | ||
| 71957 
3e162c63371a
build bit operations on word on library theory on bit operations
 haftmann parents: 
71953diff
changeset | 231 | |
| 75962 | 232 | lemma bit_numeral_even [simp]: | 
| 233 | \<open>numeral (Num.Bit0 n) = (0 :: bit)\<close> | |
| 71988 | 234 | by (simp only: Z2.bit_eq_iff even_numeral) simp | 
| 71957 
3e162c63371a
build bit operations on word on library theory on bit operations
 haftmann parents: 
71953diff
changeset | 235 | |
| 75962 | 236 | lemma bit_numeral_odd [simp]: | 
| 237 | \<open>numeral (Num.Bit1 n) = (1 :: bit)\<close> | |
| 71988 | 238 | by (simp only: Z2.bit_eq_iff odd_numeral) simp | 
| 53063 
8f7ac535892f
explicit conversion from and to bool, and into algebraic structures with 0 and 1
 haftmann parents: 
49834diff
changeset | 239 | |
| 
8f7ac535892f
explicit conversion from and to bool, and into algebraic structures with 0 and 1
 haftmann parents: 
49834diff
changeset | 240 | end |