163 definition |
163 definition |
164 word_pred :: "'a :: len0 word => 'a word" |
164 word_pred :: "'a :: len0 word => 'a word" |
165 where |
165 where |
166 "word_pred a = word_of_int (Int.pred (uint a))" |
166 "word_pred a = word_of_int (Int.pred (uint a))" |
167 |
167 |
168 constdefs |
168 definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) where |
169 udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) |
|
170 "a udvd b == EX n>=0. uint b = n * uint a" |
169 "a udvd b == EX n>=0. uint b = n * uint a" |
171 |
170 |
172 word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) |
171 definition word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) where |
173 "a <=s b == sint a <= sint b" |
172 "a <=s b == sint a <= sint b" |
174 |
173 |
175 word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50) |
174 definition word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50) where |
176 "(x <s y) == (x <=s y & x ~= y)" |
175 "(x <s y) == (x <=s y & x ~= y)" |
177 |
176 |
178 |
177 |
179 |
178 |
180 subsection "Bit-wise operations" |
179 subsection "Bit-wise operations" |
221 |
220 |
222 instance .. |
221 instance .. |
223 |
222 |
224 end |
223 end |
225 |
224 |
226 constdefs |
225 definition setBit :: "'a :: len0 word => nat => 'a word" where |
227 setBit :: "'a :: len0 word => nat => 'a word" |
|
228 "setBit w n == set_bit w n True" |
226 "setBit w n == set_bit w n True" |
229 |
227 |
230 clearBit :: "'a :: len0 word => nat => 'a word" |
228 definition clearBit :: "'a :: len0 word => nat => 'a word" where |
231 "clearBit w n == set_bit w n False" |
229 "clearBit w n == set_bit w n False" |
232 |
230 |
233 |
231 |
234 subsection "Shift operations" |
232 subsection "Shift operations" |
235 |
233 |
236 constdefs |
234 definition sshiftr1 :: "'a :: len word => 'a word" where |
237 sshiftr1 :: "'a :: len word => 'a word" |
|
238 "sshiftr1 w == word_of_int (bin_rest (sint w))" |
235 "sshiftr1 w == word_of_int (bin_rest (sint w))" |
239 |
236 |
240 bshiftr1 :: "bool => 'a :: len word => 'a word" |
237 definition bshiftr1 :: "bool => 'a :: len word => 'a word" where |
241 "bshiftr1 b w == of_bl (b # butlast (to_bl w))" |
238 "bshiftr1 b w == of_bl (b # butlast (to_bl w))" |
242 |
239 |
243 sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) |
240 definition sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) where |
244 "w >>> n == (sshiftr1 ^^ n) w" |
241 "w >>> n == (sshiftr1 ^^ n) w" |
245 |
242 |
246 mask :: "nat => 'a::len word" |
243 definition mask :: "nat => 'a::len word" where |
247 "mask n == (1 << n) - 1" |
244 "mask n == (1 << n) - 1" |
248 |
245 |
249 revcast :: "'a :: len0 word => 'b :: len0 word" |
246 definition revcast :: "'a :: len0 word => 'b :: len0 word" where |
250 "revcast w == of_bl (takefill False (len_of TYPE('b)) (to_bl w))" |
247 "revcast w == of_bl (takefill False (len_of TYPE('b)) (to_bl w))" |
251 |
248 |
252 slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" |
249 definition slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" where |
253 "slice1 n w == of_bl (takefill False n (to_bl w))" |
250 "slice1 n w == of_bl (takefill False n (to_bl w))" |
254 |
251 |
255 slice :: "nat => 'a :: len0 word => 'b :: len0 word" |
252 definition slice :: "nat => 'a :: len0 word => 'b :: len0 word" where |
256 "slice n w == slice1 (size w - n) w" |
253 "slice n w == slice1 (size w - n) w" |
257 |
254 |
258 |
255 |
259 subsection "Rotation" |
256 subsection "Rotation" |
260 |
257 |
261 constdefs |
258 definition rotater1 :: "'a list => 'a list" where |
262 rotater1 :: "'a list => 'a list" |
|
263 "rotater1 ys == |
259 "rotater1 ys == |
264 case ys of [] => [] | x # xs => last ys # butlast ys" |
260 case ys of [] => [] | x # xs => last ys # butlast ys" |
265 |
261 |
266 rotater :: "nat => 'a list => 'a list" |
262 definition rotater :: "nat => 'a list => 'a list" where |
267 "rotater n == rotater1 ^^ n" |
263 "rotater n == rotater1 ^^ n" |
268 |
264 |
269 word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" |
265 definition word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" where |
270 "word_rotr n w == of_bl (rotater n (to_bl w))" |
266 "word_rotr n w == of_bl (rotater n (to_bl w))" |
271 |
267 |
272 word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" |
268 definition word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" where |
273 "word_rotl n w == of_bl (rotate n (to_bl w))" |
269 "word_rotl n w == of_bl (rotate n (to_bl w))" |
274 |
270 |
275 word_roti :: "int => 'a :: len0 word => 'a :: len0 word" |
271 definition word_roti :: "int => 'a :: len0 word => 'a :: len0 word" where |
276 "word_roti i w == if i >= 0 then word_rotr (nat i) w |
272 "word_roti i w == if i >= 0 then word_rotr (nat i) w |
277 else word_rotl (nat (- i)) w" |
273 else word_rotl (nat (- i)) w" |
278 |
274 |
279 |
275 |
280 subsection "Split and cat operations" |
276 subsection "Split and cat operations" |
281 |
277 |
282 constdefs |
278 definition word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" where |
283 word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" |
|
284 "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))" |
279 "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))" |
285 |
280 |
286 word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" |
281 definition word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" where |
287 "word_split a == |
282 "word_split a == |
288 case bin_split (len_of TYPE ('c)) (uint a) of |
283 case bin_split (len_of TYPE ('c)) (uint a) of |
289 (u, v) => (word_of_int u, word_of_int v)" |
284 (u, v) => (word_of_int u, word_of_int v)" |
290 |
285 |
291 word_rcat :: "'a :: len0 word list => 'b :: len0 word" |
286 definition word_rcat :: "'a :: len0 word list => 'b :: len0 word" where |
292 "word_rcat ws == |
287 "word_rcat ws == |
293 word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))" |
288 word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))" |
294 |
289 |
295 word_rsplit :: "'a :: len0 word => 'b :: len word list" |
290 definition word_rsplit :: "'a :: len0 word => 'b :: len word list" where |
296 "word_rsplit w == |
291 "word_rsplit w == |
297 map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))" |
292 map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))" |
298 |
293 |
299 constdefs |
294 definition max_word :: "'a::len word" -- "Largest representable machine integer." where |
300 -- "Largest representable machine integer." |
|
301 max_word :: "'a::len word" |
|
302 "max_word \<equiv> word_of_int (2 ^ len_of TYPE('a) - 1)" |
295 "max_word \<equiv> word_of_int (2 ^ len_of TYPE('a) - 1)" |
303 |
296 |
304 consts |
297 primrec of_bool :: "bool \<Rightarrow> 'a::len word" where |
305 of_bool :: "bool \<Rightarrow> 'a::len word" |
|
306 primrec |
|
307 "of_bool False = 0" |
298 "of_bool False = 0" |
308 "of_bool True = 1" |
299 | "of_bool True = 1" |
309 |
300 |
310 |
301 |
311 lemmas of_nth_def = word_set_bits_def |
302 lemmas of_nth_def = word_set_bits_def |
312 |
303 |
313 lemmas word_size_gt_0 [iff] = |
304 lemmas word_size_gt_0 [iff] = |