src/HOL/Word/WordDefinition.thy
changeset 24374 bb0d3b49fef0
parent 24365 038b164edfef
child 24375 4aa80fadc071
equal deleted inserted replaced
24373:eb199bbbaec0 24374:bb0d3b49fef0
   173   "setBit w n == set_bit w n True"
   173   "setBit w n == set_bit w n True"
   174 
   174 
   175   clearBit :: "'a :: len0 word => nat => 'a word" 
   175   clearBit :: "'a :: len0 word => nat => 'a word" 
   176   "clearBit w n == set_bit w n False"
   176   "clearBit w n == set_bit w n False"
   177 
   177 
   178 
       
   179 subsection "Shift operations"
       
   180 
       
   181 constdefs
       
   182   shiftl1 :: "'a :: len0 word => 'a word"
       
   183   "shiftl1 w == word_of_int (uint w BIT bit.B0)"
       
   184 
       
   185   -- "shift right as unsigned or as signed, ie logical or arithmetic"
       
   186   shiftr1 :: "'a :: len0 word => 'a word"
       
   187   "shiftr1 w == word_of_int (bin_rest (uint w))"
       
   188 
       
   189   sshiftr1 :: "'a :: len word => 'a word" 
       
   190   "sshiftr1 w == word_of_int (bin_rest (sint w))"
       
   191 
       
   192   bshiftr1 :: "bool => 'a :: len word => 'a word"
       
   193   "bshiftr1 b w == of_bl (b # butlast (to_bl w))"
       
   194 
       
   195   sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55)
       
   196   "w >>> n == (sshiftr1 ^ n) w"
       
   197 
       
   198   mask :: "nat => 'a::len word"
       
   199   "mask n == (1 << n) - 1"
       
   200 
       
   201   revcast :: "'a :: len0 word => 'b :: len0 word"
       
   202   "revcast w ==  of_bl (takefill False (len_of TYPE('b)) (to_bl w))"
       
   203 
       
   204   slice1 :: "nat => 'a :: len0 word => 'b :: len0 word"
       
   205   "slice1 n w == of_bl (takefill False n (to_bl w))"
       
   206 
       
   207   slice :: "nat => 'a :: len0 word => 'b :: len0 word"
       
   208   "slice n w == slice1 (size w - n) w"
       
   209 
       
   210 
       
   211 defs (overloaded)
       
   212   shiftl_def: "(w::'a::len0 word) << n == (shiftl1 ^ n) w"
       
   213   shiftr_def: "(w::'a::len0 word) >> n == (shiftr1 ^ n) w"
       
   214 
       
   215 
       
   216 subsection "Rotation"
       
   217 
       
   218 constdefs
       
   219   rotater1 :: "'a list => 'a list"
       
   220   "rotater1 ys == 
       
   221     case ys of [] => [] | x # xs => last ys # butlast ys"
       
   222 
       
   223   rotater :: "nat => 'a list => 'a list"
       
   224   "rotater n == rotater1 ^ n"
       
   225 
       
   226   word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word"
       
   227   "word_rotr n w == of_bl (rotater n (to_bl w))"
       
   228 
       
   229   word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word"
       
   230   "word_rotl n w == of_bl (rotate n (to_bl w))"
       
   231 
       
   232   word_roti :: "int => 'a :: len0 word => 'a :: len0 word"
       
   233   "word_roti i w == if i >= 0 then word_rotr (nat i) w
       
   234                     else word_rotl (nat (- i)) w"
       
   235 
       
   236 
       
   237 subsection "Split and cat operations"
       
   238 
       
   239 constdefs
       
   240   word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word"
       
   241   "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))"
       
   242 
       
   243   word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)"
       
   244   "word_split a == 
       
   245    case bin_split (len_of TYPE ('c)) (uint a) of 
       
   246      (u, v) => (word_of_int u, word_of_int v)"
       
   247 
       
   248   word_rcat :: "'a :: len0 word list => 'b :: len0 word"
       
   249   "word_rcat ws == 
       
   250   word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))"
       
   251 
       
   252   word_rsplit :: "'a :: len0 word => 'b :: len word list"
       
   253   "word_rsplit w == 
       
   254   map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))"
       
   255 
   178 
   256 constdefs
   179 constdefs
   257   -- "Largest representable machine integer."
   180   -- "Largest representable machine integer."
   258   max_word :: "'a::len word"
   181   max_word :: "'a::len word"
   259   "max_word \<equiv> word_of_int (2^len_of TYPE('a) - 1)"
   182   "max_word \<equiv> word_of_int (2^len_of TYPE('a) - 1)"
   895 lemma ucast_down_bl': "uc = ucast ==> is_down uc ==> uc (of_bl bl) = of_bl bl"
   818 lemma ucast_down_bl': "uc = ucast ==> is_down uc ==> uc (of_bl bl) = of_bl bl"
   896   unfolding of_bl_no by clarify (erule ucast_down_no)
   819   unfolding of_bl_no by clarify (erule ucast_down_no)
   897     
   820     
   898 lemmas ucast_down_bl = ucast_down_bl' [OF refl]
   821 lemmas ucast_down_bl = ucast_down_bl' [OF refl]
   899 
   822 
   900 lemmas slice_def' = slice_def [unfolded word_size]
       
   901 lemmas test_bit_def' = word_test_bit_def [THEN meta_eq_to_obj_eq, THEN fun_cong]
   823 lemmas test_bit_def' = word_test_bit_def [THEN meta_eq_to_obj_eq, THEN fun_cong]
   902 
   824 
   903 lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def
   825 lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def
   904 lemmas word_log_bin_defs = word_log_defs
   826 lemmas word_log_bin_defs = word_log_defs
   905 
   827