| 24333 |      1 | (* 
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Jeremy Dawson and Gerwin Klein, NICTA
 | 
|  |      4 | 
 | 
|  |      5 |   contains theorems to do with shifting, rotating, splitting words
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
| 24350 |      8 | header {* Shifting, Rotating, and Splitting Words *}
 | 
| 24465 |      9 | 
 | 
|  |     10 | theory WordShift imports WordBitwise begin
 | 
| 24333 |     11 | 
 | 
| 24350 |     12 | subsection "Bit shifting"
 | 
| 24333 |     13 | 
 | 
|  |     14 | lemma shiftl1_number [simp] :
 | 
|  |     15 |   "shiftl1 (number_of w) = number_of (w BIT bit.B0)"
 | 
|  |     16 |   apply (unfold shiftl1_def word_number_of_def)
 | 
|  |     17 |   apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm)
 | 
|  |     18 |   apply (subst refl [THEN bintrunc_BIT_I, symmetric])
 | 
|  |     19 |   apply (subst bintrunc_bintrunc_min)
 | 
|  |     20 |   apply simp
 | 
|  |     21 |   done
 | 
|  |     22 | 
 | 
|  |     23 | lemma shiftl1_0 [simp] : "shiftl1 0 = 0"
 | 
|  |     24 |   unfolding word_0_no shiftl1_number by auto
 | 
|  |     25 | 
 | 
|  |     26 | lemmas shiftl1_def_u = shiftl1_def [folded word_number_of_def]
 | 
|  |     27 | 
 | 
|  |     28 | lemma shiftl1_def_s: "shiftl1 w = number_of (sint w BIT bit.B0)"
 | 
|  |     29 |   by (rule trans [OF _ shiftl1_number]) simp
 | 
|  |     30 | 
 | 
|  |     31 | lemma shiftr1_0 [simp] : "shiftr1 0 = 0"
 | 
|  |     32 |   unfolding shiftr1_def 
 | 
|  |     33 |   by simp (simp add: word_0_wi)
 | 
|  |     34 | 
 | 
|  |     35 | lemma sshiftr1_0 [simp] : "sshiftr1 0 = 0"
 | 
|  |     36 |   apply (unfold sshiftr1_def)
 | 
|  |     37 |   apply simp
 | 
|  |     38 |   apply (simp add : word_0_wi)
 | 
|  |     39 |   done
 | 
|  |     40 | 
 | 
|  |     41 | lemma sshiftr1_n1 [simp] : "sshiftr1 -1 = -1"
 | 
|  |     42 |   unfolding sshiftr1_def by auto
 | 
|  |     43 | 
 | 
| 24465 |     44 | lemma shiftl_0 [simp] : "(0::'a::len0 word) << n = 0"
 | 
| 24333 |     45 |   unfolding shiftl_def by (induct n) auto
 | 
|  |     46 | 
 | 
| 24465 |     47 | lemma shiftr_0 [simp] : "(0::'a::len0 word) >> n = 0"
 | 
| 24333 |     48 |   unfolding shiftr_def by (induct n) auto
 | 
|  |     49 | 
 | 
|  |     50 | lemma sshiftr_0 [simp] : "0 >>> n = 0"
 | 
|  |     51 |   unfolding sshiftr_def by (induct n) auto
 | 
|  |     52 | 
 | 
|  |     53 | lemma sshiftr_n1 [simp] : "-1 >>> n = -1"
 | 
|  |     54 |   unfolding sshiftr_def by (induct n) auto
 | 
|  |     55 | 
 | 
|  |     56 | lemma nth_shiftl1: "shiftl1 w !! n = (n < size w & n > 0 & w !! (n - 1))"
 | 
|  |     57 |   apply (unfold shiftl1_def word_test_bit_def)
 | 
|  |     58 |   apply (simp add: nth_bintr word_ubin.eq_norm word_size)
 | 
|  |     59 |   apply (cases n)
 | 
|  |     60 |    apply auto
 | 
|  |     61 |   done
 | 
|  |     62 | 
 | 
|  |     63 | lemma nth_shiftl' [rule_format]:
 | 
| 24465 |     64 |   "ALL n. ((w::'a::len0 word) << m) !! n = (n < size w & n >= m & w !! (n - m))"
 | 
| 24333 |     65 |   apply (unfold shiftl_def)
 | 
|  |     66 |   apply (induct "m")
 | 
|  |     67 |    apply (force elim!: test_bit_size)
 | 
|  |     68 |   apply (clarsimp simp add : nth_shiftl1 word_size)
 | 
|  |     69 |   apply arith
 | 
|  |     70 |   done
 | 
|  |     71 | 
 | 
|  |     72 | lemmas nth_shiftl = nth_shiftl' [unfolded word_size] 
 | 
|  |     73 | 
 | 
|  |     74 | lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n"
 | 
|  |     75 |   apply (unfold shiftr1_def word_test_bit_def)
 | 
|  |     76 |   apply (simp add: nth_bintr word_ubin.eq_norm)
 | 
|  |     77 |   apply safe
 | 
|  |     78 |   apply (drule bin_nth.Suc [THEN iffD2, THEN bin_nth_uint_imp])
 | 
|  |     79 |   apply simp
 | 
|  |     80 |   done
 | 
|  |     81 | 
 | 
|  |     82 | lemma nth_shiftr: 
 | 
| 24465 |     83 |   "\<And>n. ((w::'a::len0 word) >> m) !! n = w !! (n + m)"
 | 
| 24333 |     84 |   apply (unfold shiftr_def)
 | 
|  |     85 |   apply (induct "m")
 | 
|  |     86 |    apply (auto simp add : nth_shiftr1)
 | 
|  |     87 |   done
 | 
|  |     88 |    
 | 
|  |     89 | (* see paper page 10, (1), (2), shiftr1_def is of the form of (1),
 | 
|  |     90 |   where f (ie bin_rest) takes normal arguments to normal results,
 | 
|  |     91 |   thus we get (2) from (1) *)
 | 
|  |     92 | 
 | 
|  |     93 | lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" 
 | 
|  |     94 |   apply (unfold shiftr1_def word_ubin.eq_norm bin_rest_trunc_i)
 | 
|  |     95 |   apply (subst bintr_uint [symmetric, OF order_refl])
 | 
|  |     96 |   apply (simp only : bintrunc_bintrunc_l)
 | 
|  |     97 |   apply simp 
 | 
|  |     98 |   done
 | 
|  |     99 | 
 | 
|  |    100 | lemma nth_sshiftr1: 
 | 
|  |    101 |   "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)"
 | 
|  |    102 |   apply (unfold sshiftr1_def word_test_bit_def)
 | 
|  |    103 |   apply (simp add: nth_bintr word_ubin.eq_norm
 | 
|  |    104 |                    bin_nth.Suc [symmetric] word_size 
 | 
|  |    105 |              del: bin_nth.simps)
 | 
|  |    106 |   apply (simp add: nth_bintr uint_sint del : bin_nth.simps)
 | 
|  |    107 |   apply (auto simp add: bin_nth_sint)
 | 
|  |    108 |   done
 | 
|  |    109 | 
 | 
|  |    110 | lemma nth_sshiftr [rule_format] : 
 | 
|  |    111 |   "ALL n. sshiftr w m !! n = (n < size w & 
 | 
|  |    112 |     (if n + m >= size w then w !! (size w - 1) else w !! (n + m)))"
 | 
|  |    113 |   apply (unfold sshiftr_def)
 | 
|  |    114 |   apply (induct_tac "m")
 | 
|  |    115 |    apply (simp add: test_bit_bl)
 | 
|  |    116 |   apply (clarsimp simp add: nth_sshiftr1 word_size)
 | 
|  |    117 |   apply safe
 | 
|  |    118 |        apply arith
 | 
|  |    119 |       apply arith
 | 
|  |    120 |      apply (erule thin_rl)
 | 
|  |    121 |      apply (case_tac n)
 | 
|  |    122 |       apply safe
 | 
|  |    123 |       apply simp
 | 
|  |    124 |      apply simp
 | 
|  |    125 |     apply (erule thin_rl)
 | 
|  |    126 |     apply (case_tac n)
 | 
|  |    127 |      apply safe
 | 
|  |    128 |      apply simp
 | 
|  |    129 |     apply simp
 | 
|  |    130 |    apply arith+
 | 
|  |    131 |   done
 | 
|  |    132 |     
 | 
|  |    133 | lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2"
 | 
|  |    134 |   apply (unfold shiftr1_def bin_rest_div)
 | 
|  |    135 |   apply (rule word_uint.Abs_inverse)
 | 
|  |    136 |   apply (simp add: uints_num pos_imp_zdiv_nonneg_iff)
 | 
|  |    137 |   apply (rule xtr7)
 | 
|  |    138 |    prefer 2
 | 
|  |    139 |    apply (rule zdiv_le_dividend)
 | 
|  |    140 |     apply auto
 | 
|  |    141 |   done
 | 
|  |    142 | 
 | 
|  |    143 | lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2"
 | 
|  |    144 |   apply (unfold sshiftr1_def bin_rest_div [symmetric])
 | 
|  |    145 |   apply (simp add: word_sbin.eq_norm)
 | 
|  |    146 |   apply (rule trans)
 | 
|  |    147 |    defer
 | 
|  |    148 |    apply (subst word_sbin.norm_Rep [symmetric])
 | 
|  |    149 |    apply (rule refl)
 | 
|  |    150 |   apply (subst word_sbin.norm_Rep [symmetric])
 | 
|  |    151 |   apply (unfold One_nat_def)
 | 
|  |    152 |   apply (rule sbintrunc_rest)
 | 
|  |    153 |   done
 | 
|  |    154 | 
 | 
|  |    155 | lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n"
 | 
|  |    156 |   apply (unfold shiftr_def)
 | 
|  |    157 |   apply (induct "n")
 | 
|  |    158 |    apply simp
 | 
|  |    159 |   apply (simp add: shiftr1_div_2 mult_commute
 | 
|  |    160 |                    zdiv_zmult2_eq [symmetric])
 | 
|  |    161 |   done
 | 
|  |    162 | 
 | 
|  |    163 | lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n"
 | 
|  |    164 |   apply (unfold sshiftr_def)
 | 
|  |    165 |   apply (induct "n")
 | 
|  |    166 |    apply simp
 | 
|  |    167 |   apply (simp add: sshiftr1_div_2 mult_commute
 | 
|  |    168 |                    zdiv_zmult2_eq [symmetric])
 | 
|  |    169 |   done
 | 
|  |    170 | 
 | 
| 24350 |    171 | subsubsection "shift functions in terms of lists of bools"
 | 
| 24333 |    172 | 
 | 
|  |    173 | lemmas bshiftr1_no_bin [simp] = 
 | 
|  |    174 |   bshiftr1_def [where w="number_of ?w", unfolded to_bl_no_bin]
 | 
|  |    175 | 
 | 
|  |    176 | lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)"
 | 
|  |    177 |   unfolding bshiftr1_def by (rule word_bl.Abs_inverse) simp
 | 
|  |    178 | 
 | 
|  |    179 | lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])"
 | 
|  |    180 |   unfolding uint_bl of_bl_no 
 | 
|  |    181 |   by (simp add: bl_to_bin_aux_append bl_to_bin_def)
 | 
|  |    182 | 
 | 
|  |    183 | lemmas shiftl1_bl = shiftl1_of_bl 
 | 
| 24465 |    184 |   [where bl = "to_bl (?w :: ?'a :: len0 word)", simplified]
 | 
| 24333 |    185 | 
 | 
|  |    186 | lemma bl_shiftl1:
 | 
| 24465 |    187 |   "to_bl (shiftl1 (w :: 'a :: len word)) = tl (to_bl w) @ [False]"
 | 
|  |    188 |   apply (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons')
 | 
|  |    189 |   apply (fast intro!: Suc_leI)
 | 
|  |    190 |   done
 | 
| 24333 |    191 | 
 | 
|  |    192 | lemma shiftr1_bl: "shiftr1 w = of_bl (butlast (to_bl w))"
 | 
|  |    193 |   apply (unfold shiftr1_def uint_bl of_bl_def)
 | 
|  |    194 |   apply (simp add: butlast_rest_bin word_size)
 | 
|  |    195 |   apply (simp add: bin_rest_trunc [symmetric, unfolded One_nat_def])
 | 
|  |    196 |   done
 | 
|  |    197 | 
 | 
|  |    198 | lemma bl_shiftr1: 
 | 
| 24465 |    199 |   "to_bl (shiftr1 (w :: 'a :: len word)) = False # butlast (to_bl w)"
 | 
| 24333 |    200 |   unfolding shiftr1_bl
 | 
| 24465 |    201 |   by (simp add : word_rep_drop len_gt_0 [THEN Suc_leI])
 | 
| 24333 |    202 | 
 | 
|  |    203 | 
 | 
| 24465 |    204 | (* relate the two above : TODO - remove the :: len restriction on
 | 
| 24333 |    205 |   this theorem and others depending on it *)
 | 
|  |    206 | lemma shiftl1_rev: 
 | 
| 24465 |    207 |   "shiftl1 (w :: 'a :: len word) = word_reverse (shiftr1 (word_reverse w))"
 | 
| 24333 |    208 |   apply (unfold word_reverse_def)
 | 
|  |    209 |   apply (rule word_bl.Rep_inverse' [symmetric])
 | 
|  |    210 |   apply (simp add: bl_shiftl1 bl_shiftr1 word_bl.Abs_inverse)
 | 
|  |    211 |   apply (cases "to_bl w")
 | 
|  |    212 |    apply auto
 | 
|  |    213 |   done
 | 
|  |    214 | 
 | 
|  |    215 | lemma shiftl_rev: 
 | 
| 24465 |    216 |   "shiftl (w :: 'a :: len word) n = word_reverse (shiftr (word_reverse w) n)"
 | 
| 24333 |    217 |   apply (unfold shiftl_def shiftr_def)
 | 
|  |    218 |   apply (induct "n")
 | 
|  |    219 |    apply (auto simp add : shiftl1_rev)
 | 
|  |    220 |   done
 | 
|  |    221 | 
 | 
|  |    222 | lemmas rev_shiftl =
 | 
|  |    223 |   shiftl_rev [where w = "word_reverse ?w1", simplified, standard]
 | 
|  |    224 | 
 | 
|  |    225 | lemmas shiftr_rev = rev_shiftl [THEN word_rev_gal', standard]
 | 
|  |    226 | lemmas rev_shiftr = shiftl_rev [THEN word_rev_gal', standard]
 | 
|  |    227 | 
 | 
|  |    228 | lemma bl_sshiftr1:
 | 
| 24465 |    229 |   "to_bl (sshiftr1 (w :: 'a :: len word)) = hd (to_bl w) # butlast (to_bl w)"
 | 
| 24333 |    230 |   apply (unfold sshiftr1_def uint_bl word_size)
 | 
|  |    231 |   apply (simp add: butlast_rest_bin word_ubin.eq_norm)
 | 
|  |    232 |   apply (simp add: sint_uint)
 | 
|  |    233 |   apply (rule nth_equalityI)
 | 
|  |    234 |    apply clarsimp
 | 
|  |    235 |   apply clarsimp
 | 
|  |    236 |   apply (case_tac i)
 | 
|  |    237 |    apply (simp_all add: hd_conv_nth length_0_conv [symmetric]
 | 
|  |    238 |                         nth_bin_to_bl bin_nth.Suc [symmetric] 
 | 
|  |    239 |                         nth_sbintr 
 | 
|  |    240 |                    del: bin_nth.Suc)
 | 
| 24465 |    241 |    apply force
 | 
| 24333 |    242 |   apply (rule impI)
 | 
|  |    243 |   apply (rule_tac f = "bin_nth (uint w)" in arg_cong)
 | 
|  |    244 |   apply simp
 | 
|  |    245 |   done
 | 
|  |    246 | 
 | 
|  |    247 | lemma drop_shiftr: 
 | 
| 24465 |    248 |   "drop n (to_bl ((w :: 'a :: len word) >> n)) = take (size w - n) (to_bl w)" 
 | 
| 24333 |    249 |   apply (unfold shiftr_def)
 | 
|  |    250 |   apply (induct n)
 | 
|  |    251 |    prefer 2
 | 
|  |    252 |    apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric])
 | 
|  |    253 |    apply (rule butlast_take [THEN trans])
 | 
|  |    254 |   apply (auto simp: word_size)
 | 
|  |    255 |   done
 | 
|  |    256 | 
 | 
|  |    257 | lemma drop_sshiftr: 
 | 
| 24465 |    258 |   "drop n (to_bl ((w :: 'a :: len word) >>> n)) = take (size w - n) (to_bl w)"
 | 
| 24333 |    259 |   apply (unfold sshiftr_def)
 | 
|  |    260 |   apply (induct n)
 | 
|  |    261 |    prefer 2
 | 
|  |    262 |    apply (simp add: drop_Suc bl_sshiftr1 butlast_drop [symmetric])
 | 
|  |    263 |    apply (rule butlast_take [THEN trans])
 | 
|  |    264 |   apply (auto simp: word_size)
 | 
|  |    265 |   done
 | 
|  |    266 | 
 | 
|  |    267 | lemma take_shiftr [rule_format] :
 | 
| 24465 |    268 |   "n <= size (w :: 'a :: len word) --> take n (to_bl (w >> n)) = 
 | 
| 24333 |    269 |     replicate n False" 
 | 
|  |    270 |   apply (unfold shiftr_def)
 | 
|  |    271 |   apply (induct n)
 | 
|  |    272 |    prefer 2
 | 
|  |    273 |    apply (simp add: bl_shiftr1)
 | 
|  |    274 |    apply (rule impI)
 | 
|  |    275 |    apply (rule take_butlast [THEN trans])
 | 
|  |    276 |   apply (auto simp: word_size)
 | 
|  |    277 |   done
 | 
|  |    278 | 
 | 
|  |    279 | lemma take_sshiftr' [rule_format] :
 | 
| 24465 |    280 |   "n <= size (w :: 'a :: len word) --> hd (to_bl (w >>> n)) = hd (to_bl w) & 
 | 
| 24333 |    281 |     take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" 
 | 
|  |    282 |   apply (unfold sshiftr_def)
 | 
|  |    283 |   apply (induct n)
 | 
|  |    284 |    prefer 2
 | 
|  |    285 |    apply (simp add: bl_sshiftr1)
 | 
|  |    286 |    apply (rule impI)
 | 
|  |    287 |    apply (rule take_butlast [THEN trans])
 | 
|  |    288 |   apply (auto simp: word_size)
 | 
|  |    289 |   done
 | 
|  |    290 | 
 | 
|  |    291 | lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1, standard]
 | 
|  |    292 | lemmas take_sshiftr = take_sshiftr' [THEN conjunct2, standard]
 | 
|  |    293 | 
 | 
|  |    294 | lemma atd_lem: "take n xs = t ==> drop n xs = d ==> xs = t @ d"
 | 
|  |    295 |   by (auto intro: append_take_drop_id [symmetric])
 | 
|  |    296 | 
 | 
|  |    297 | lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr]
 | 
|  |    298 | lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr]
 | 
|  |    299 | 
 | 
|  |    300 | lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)"
 | 
|  |    301 |   unfolding shiftl_def
 | 
|  |    302 |   by (induct n) (auto simp: shiftl1_of_bl replicate_app_Cons_same)
 | 
|  |    303 | 
 | 
|  |    304 | lemmas shiftl_bl =
 | 
| 24465 |    305 |   shiftl_of_bl [where bl = "to_bl (?w :: ?'a :: len0 word)", simplified]
 | 
| 24333 |    306 | 
 | 
|  |    307 | lemmas shiftl_number [simp] = shiftl_def [where w="number_of ?w"]
 | 
|  |    308 | 
 | 
|  |    309 | lemma bl_shiftl:
 | 
|  |    310 |   "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False"
 | 
|  |    311 |   by (simp add: shiftl_bl word_rep_drop word_size min_def)
 | 
|  |    312 | 
 | 
|  |    313 | lemma shiftl_zero_size: 
 | 
| 24465 |    314 |   fixes x :: "'a::len0 word"
 | 
| 24333 |    315 |   shows "size x <= n ==> x << n = 0"
 | 
|  |    316 |   apply (unfold word_size)
 | 
|  |    317 |   apply (rule word_eqI)
 | 
|  |    318 |   apply (clarsimp simp add: shiftl_bl word_size test_bit_of_bl nth_append)
 | 
|  |    319 |   done
 | 
|  |    320 | 
 | 
| 24465 |    321 | (* note - the following results use 'a :: len word < number_ring *)
 | 
| 24333 |    322 | 
 | 
| 24465 |    323 | lemma shiftl1_2t: "shiftl1 (w :: 'a :: len word) = 2 * w"
 | 
| 24333 |    324 |   apply (simp add: shiftl1_def_u)
 | 
|  |    325 |   apply (simp only:  double_number_of_BIT [symmetric])
 | 
|  |    326 |   apply simp
 | 
|  |    327 |   done
 | 
|  |    328 | 
 | 
| 24465 |    329 | lemma shiftl1_p: "shiftl1 (w :: 'a :: len word) = w + w"
 | 
| 24333 |    330 |   apply (simp add: shiftl1_def_u)
 | 
|  |    331 |   apply (simp only: double_number_of_BIT [symmetric])
 | 
|  |    332 |   apply simp
 | 
|  |    333 |   done
 | 
|  |    334 | 
 | 
| 24465 |    335 | lemma shiftl_t2n: "shiftl (w :: 'a :: len word) n = 2 ^ n * w"
 | 
| 24333 |    336 |   unfolding shiftl_def 
 | 
|  |    337 |   by (induct n) (auto simp: shiftl1_2t power_Suc)
 | 
|  |    338 | 
 | 
|  |    339 | lemma shiftr1_bintr [simp]:
 | 
| 24465 |    340 |   "(shiftr1 (number_of w) :: 'a :: len0 word) = 
 | 
|  |    341 |     number_of (bin_rest (bintrunc (len_of TYPE ('a)) w))" 
 | 
| 24333 |    342 |   unfolding shiftr1_def word_number_of_def
 | 
|  |    343 |   by (simp add : word_ubin.eq_norm)
 | 
|  |    344 | 
 | 
|  |    345 | lemma sshiftr1_sbintr [simp] :
 | 
| 24465 |    346 |   "(sshiftr1 (number_of w) :: 'a :: len word) = 
 | 
|  |    347 |     number_of (bin_rest (sbintrunc (len_of TYPE ('a) - 1) w))" 
 | 
| 24333 |    348 |   unfolding sshiftr1_def word_number_of_def
 | 
|  |    349 |   by (simp add : word_sbin.eq_norm)
 | 
|  |    350 | 
 | 
|  |    351 | lemma shiftr_no': 
 | 
|  |    352 |   "w = number_of bin ==> 
 | 
| 24465 |    353 |   (w::'a::len0 word) >> n = number_of ((bin_rest ^ n) (bintrunc (size w) bin))"
 | 
| 24333 |    354 |   apply clarsimp
 | 
|  |    355 |   apply (rule word_eqI)
 | 
|  |    356 |   apply (auto simp: nth_shiftr nth_rest_power_bin nth_bintr word_size)
 | 
|  |    357 |   done
 | 
|  |    358 | 
 | 
|  |    359 | lemma sshiftr_no': 
 | 
|  |    360 |   "w = number_of bin ==> w >>> n = number_of ((bin_rest ^ n) 
 | 
|  |    361 |     (sbintrunc (size w - 1) bin))"
 | 
|  |    362 |   apply clarsimp
 | 
|  |    363 |   apply (rule word_eqI)
 | 
|  |    364 |   apply (auto simp: nth_sshiftr nth_rest_power_bin nth_sbintr word_size)
 | 
| 24465 |    365 |    apply (subgoal_tac "na + n = len_of TYPE('a) - Suc 0", simp, simp)+
 | 
| 24333 |    366 |   done
 | 
|  |    367 | 
 | 
|  |    368 | lemmas sshiftr_no [simp] = 
 | 
|  |    369 |   sshiftr_no' [where w = "number_of ?w", OF refl, unfolded word_size]
 | 
|  |    370 | 
 | 
|  |    371 | lemmas shiftr_no [simp] = 
 | 
|  |    372 |   shiftr_no' [where w = "number_of ?w", OF refl, unfolded word_size]
 | 
|  |    373 | 
 | 
|  |    374 | lemma shiftr1_bl_of': 
 | 
|  |    375 |   "us = shiftr1 (of_bl bl) ==> length bl <= size us ==> 
 | 
|  |    376 |     us = of_bl (butlast bl)"
 | 
|  |    377 |   by (clarsimp simp: shiftr1_def of_bl_def word_size butlast_rest_bl2bin 
 | 
|  |    378 |                      word_ubin.eq_norm trunc_bl2bin)
 | 
|  |    379 | 
 | 
|  |    380 | lemmas shiftr1_bl_of = refl [THEN shiftr1_bl_of', unfolded word_size]
 | 
|  |    381 | 
 | 
|  |    382 | lemma shiftr_bl_of' [rule_format]: 
 | 
|  |    383 |   "us = of_bl bl >> n ==> length bl <= size us --> 
 | 
|  |    384 |    us = of_bl (take (length bl - n) bl)"
 | 
|  |    385 |   apply (unfold shiftr_def)
 | 
|  |    386 |   apply hypsubst
 | 
|  |    387 |   apply (unfold word_size)
 | 
|  |    388 |   apply (induct n)
 | 
|  |    389 |    apply clarsimp
 | 
|  |    390 |   apply clarsimp
 | 
|  |    391 |   apply (subst shiftr1_bl_of)
 | 
|  |    392 |    apply simp
 | 
|  |    393 |   apply (simp add: butlast_take)
 | 
|  |    394 |   done
 | 
|  |    395 | 
 | 
|  |    396 | lemmas shiftr_bl_of = refl [THEN shiftr_bl_of', unfolded word_size]
 | 
|  |    397 | 
 | 
|  |    398 | lemmas shiftr_bl = word_bl.Rep' [THEN eq_imp_le, THEN shiftr_bl_of,
 | 
|  |    399 |   simplified word_size, simplified, THEN eq_reflection, standard]
 | 
|  |    400 | 
 | 
| 24465 |    401 | lemma msb_shift': "msb (w::'a::len word) <-> (w >> (size w - 1)) ~= 0"
 | 
| 24333 |    402 |   apply (unfold shiftr_bl word_msb_alt)
 | 
|  |    403 |   apply (simp add: word_size Suc_le_eq take_Suc)
 | 
|  |    404 |   apply (cases "hd (to_bl w)")
 | 
|  |    405 |    apply (auto simp: word_1_bl word_0_bl 
 | 
|  |    406 |                      of_bl_rep_False [where n=1 and bs="[]", simplified])
 | 
|  |    407 |   done
 | 
|  |    408 | 
 | 
|  |    409 | lemmas msb_shift = msb_shift' [unfolded word_size]
 | 
|  |    410 | 
 | 
|  |    411 | lemma align_lem_or [rule_format] :
 | 
|  |    412 |   "ALL x m. length x = n + m --> length y = n + m --> 
 | 
|  |    413 |     drop m x = replicate n False --> take m y = replicate m False --> 
 | 
|  |    414 |     app2 op | x y = take m x @ drop m y"
 | 
|  |    415 |   apply (induct_tac y)
 | 
|  |    416 |    apply force
 | 
|  |    417 |   apply clarsimp
 | 
|  |    418 |   apply (case_tac x, force)
 | 
|  |    419 |   apply (case_tac m, auto)
 | 
|  |    420 |   apply (drule sym)
 | 
|  |    421 |   apply auto
 | 
|  |    422 |   apply (induct_tac list, auto)
 | 
|  |    423 |   done
 | 
|  |    424 | 
 | 
|  |    425 | lemma align_lem_and [rule_format] :
 | 
|  |    426 |   "ALL x m. length x = n + m --> length y = n + m --> 
 | 
|  |    427 |     drop m x = replicate n False --> take m y = replicate m False --> 
 | 
|  |    428 |     app2 op & x y = replicate (n + m) False"
 | 
|  |    429 |   apply (induct_tac y)
 | 
|  |    430 |    apply force
 | 
|  |    431 |   apply clarsimp
 | 
|  |    432 |   apply (case_tac x, force)
 | 
|  |    433 |   apply (case_tac m, auto)
 | 
|  |    434 |   apply (drule sym)
 | 
|  |    435 |   apply auto
 | 
|  |    436 |   apply (induct_tac list, auto)
 | 
|  |    437 |   done
 | 
|  |    438 | 
 | 
|  |    439 | lemma aligned_bl_add_size':
 | 
|  |    440 |   "size x - n = m ==> n <= size x ==> drop m (to_bl x) = replicate n False ==>
 | 
|  |    441 |     take m (to_bl y) = replicate m False ==> 
 | 
|  |    442 |     to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)"
 | 
|  |    443 |   apply (subgoal_tac "x AND y = 0")
 | 
|  |    444 |    prefer 2
 | 
|  |    445 |    apply (rule word_bl.Rep_eqD)
 | 
|  |    446 |    apply (simp add: bl_word_and to_bl_0)
 | 
|  |    447 |    apply (rule align_lem_and [THEN trans])
 | 
|  |    448 |        apply (simp_all add: word_size)[5]
 | 
|  |    449 |    apply (rule_tac f = "%n. replicate n False" in arg_cong)
 | 
|  |    450 |    apply simp
 | 
|  |    451 |   apply (subst word_plus_and_or [symmetric])
 | 
|  |    452 |   apply (simp add : bl_word_or)
 | 
|  |    453 |   apply (rule align_lem_or)
 | 
|  |    454 |      apply (simp_all add: word_size)
 | 
|  |    455 |   done
 | 
|  |    456 | 
 | 
|  |    457 | lemmas aligned_bl_add_size = refl [THEN aligned_bl_add_size']
 | 
|  |    458 | 
 | 
| 24350 |    459 | subsubsection "Mask"
 | 
| 24333 |    460 | 
 | 
|  |    461 | lemma nth_mask': "m = mask n ==> test_bit m i = (i < n & i < size m)"
 | 
|  |    462 |   apply (unfold mask_def test_bit_bl)
 | 
|  |    463 |   apply (simp only: word_1_bl [symmetric] shiftl_of_bl)
 | 
|  |    464 |   apply (clarsimp simp add: word_size)
 | 
|  |    465 |   apply (simp only: of_bl_no mask_lem number_of_succ add_diff_cancel2)
 | 
|  |    466 |   apply (fold of_bl_no)
 | 
|  |    467 |   apply (simp add: word_1_bl)
 | 
|  |    468 |   apply (rule test_bit_of_bl [THEN trans, unfolded test_bit_bl word_size])
 | 
|  |    469 |   apply auto
 | 
|  |    470 |   done
 | 
|  |    471 | 
 | 
|  |    472 | lemmas nth_mask [simp] = refl [THEN nth_mask']
 | 
|  |    473 | 
 | 
|  |    474 | lemma mask_bl: "mask n = of_bl (replicate n True)"
 | 
|  |    475 |   by (auto simp add : test_bit_of_bl word_size intro: word_eqI)
 | 
|  |    476 | 
 | 
|  |    477 | lemma mask_bin: "mask n = number_of (bintrunc n Numeral.Min)"
 | 
|  |    478 |   by (auto simp add: nth_bintr word_size intro: word_eqI)
 | 
|  |    479 | 
 | 
|  |    480 | lemma and_mask_bintr: "w AND mask n = number_of (bintrunc n (uint w))"
 | 
|  |    481 |   apply (rule word_eqI)
 | 
|  |    482 |   apply (simp add: nth_bintr word_size word_ops_nth_size)
 | 
|  |    483 |   apply (auto simp add: test_bit_bin)
 | 
|  |    484 |   done
 | 
|  |    485 | 
 | 
|  |    486 | lemma and_mask_no: "number_of i AND mask n = number_of (bintrunc n i)" 
 | 
|  |    487 |   by (auto simp add : nth_bintr word_size word_ops_nth_size intro: word_eqI)
 | 
|  |    488 | 
 | 
|  |    489 | lemmas and_mask_wi = and_mask_no [unfolded word_number_of_def] 
 | 
|  |    490 | 
 | 
|  |    491 | lemma bl_and_mask:
 | 
| 24465 |    492 |   "to_bl (w AND mask n :: 'a :: len word) = 
 | 
|  |    493 |     replicate (len_of TYPE('a) - n) False @ 
 | 
|  |    494 |     drop (len_of TYPE('a) - n) (to_bl w)"
 | 
| 24333 |    495 |   apply (rule nth_equalityI)
 | 
|  |    496 |    apply simp
 | 
|  |    497 |   apply (clarsimp simp add: to_bl_nth word_size)
 | 
|  |    498 |   apply (simp add: word_size word_ops_nth_size)
 | 
|  |    499 |   apply (auto simp add: word_size test_bit_bl nth_append nth_rev)
 | 
|  |    500 |   done
 | 
|  |    501 | 
 | 
|  |    502 | lemmas and_mask_mod_2p = 
 | 
|  |    503 |   and_mask_bintr [unfolded word_number_of_alt no_bintr_alt]
 | 
|  |    504 | 
 | 
|  |    505 | lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n"
 | 
|  |    506 |   apply (simp add : and_mask_bintr no_bintr_alt)
 | 
|  |    507 |   apply (rule xtr8)
 | 
|  |    508 |    prefer 2
 | 
|  |    509 |    apply (rule pos_mod_bound)
 | 
|  |    510 |   apply auto
 | 
|  |    511 |   done
 | 
|  |    512 | 
 | 
|  |    513 | lemmas eq_mod_iff = trans [symmetric, OF int_mod_lem eq_sym_conv]
 | 
|  |    514 | 
 | 
|  |    515 | lemma mask_eq_iff: "(w AND mask n) = w <-> uint w < 2 ^ n"
 | 
|  |    516 |   apply (simp add: and_mask_bintr word_number_of_def)
 | 
|  |    517 |   apply (simp add: word_ubin.inverse_norm)
 | 
|  |    518 |   apply (simp add: eq_mod_iff bintrunc_mod2p min_def)
 | 
|  |    519 |   apply (fast intro!: lt2p_lem)
 | 
|  |    520 |   done
 | 
|  |    521 | 
 | 
|  |    522 | lemma and_mask_dvd: "2 ^ n dvd uint w = (w AND mask n = 0)"
 | 
|  |    523 |   apply (simp add: zdvd_iff_zmod_eq_0 and_mask_mod_2p)
 | 
|  |    524 |   apply (simp add: word_uint.norm_eq_iff [symmetric] word_of_int_homs)
 | 
|  |    525 |   apply (subst word_uint.norm_Rep [symmetric])
 | 
|  |    526 |   apply (simp only: bintrunc_bintrunc_min bintrunc_mod2p [symmetric] min_def)
 | 
|  |    527 |   apply auto
 | 
|  |    528 |   done
 | 
|  |    529 | 
 | 
|  |    530 | lemma and_mask_dvd_nat: "2 ^ n dvd unat w = (w AND mask n = 0)"
 | 
|  |    531 |   apply (unfold unat_def)
 | 
|  |    532 |   apply (rule trans [OF _ and_mask_dvd])
 | 
|  |    533 |   apply (unfold dvd_def) 
 | 
|  |    534 |   apply auto 
 | 
|  |    535 |   apply (drule uint_ge_0 [THEN nat_int.Abs_inverse' [simplified], symmetric])
 | 
|  |    536 |   apply (simp add : int_mult int_power)
 | 
|  |    537 |   apply (simp add : nat_mult_distrib nat_power_eq)
 | 
|  |    538 |   done
 | 
|  |    539 | 
 | 
|  |    540 | lemma word_2p_lem: 
 | 
| 24465 |    541 |   "n < size w ==> w < 2 ^ n = (uint (w :: 'a :: len word) < 2 ^ n)"
 | 
| 24333 |    542 |   apply (unfold word_size word_less_alt word_number_of_alt)
 | 
|  |    543 |   apply (clarsimp simp add: word_of_int_power_hom word_uint.eq_norm 
 | 
|  |    544 |                             int_mod_eq'
 | 
|  |    545 |                   simp del: word_of_int_bin)
 | 
|  |    546 |   done
 | 
|  |    547 | 
 | 
| 24465 |    548 | lemma less_mask_eq: "x < 2 ^ n ==> x AND mask n = (x :: 'a :: len word)"
 | 
| 24333 |    549 |   apply (unfold word_less_alt word_number_of_alt)
 | 
|  |    550 |   apply (clarsimp simp add: and_mask_mod_2p word_of_int_power_hom 
 | 
|  |    551 |                             word_uint.eq_norm
 | 
|  |    552 |                   simp del: word_of_int_bin)
 | 
|  |    553 |   apply (drule xtr8 [rotated])
 | 
|  |    554 |   apply (rule int_mod_le)
 | 
|  |    555 |   apply (auto simp add : mod_pos_pos_trivial)
 | 
|  |    556 |   done
 | 
|  |    557 | 
 | 
|  |    558 | lemmas mask_eq_iff_w2p =
 | 
|  |    559 |   trans [OF mask_eq_iff word_2p_lem [symmetric], standard]
 | 
|  |    560 | 
 | 
|  |    561 | lemmas and_mask_less' = 
 | 
|  |    562 |   iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size, standard]
 | 
|  |    563 | 
 | 
|  |    564 | lemma and_mask_less_size: "n < size x ==> x AND mask n < 2^n"
 | 
|  |    565 |   unfolding word_size by (erule and_mask_less')
 | 
|  |    566 | 
 | 
|  |    567 | lemma word_mod_2p_is_mask':
 | 
| 24465 |    568 |   "c = 2 ^ n ==> c > 0 ==> x mod c = (x :: 'a :: len word) AND mask n" 
 | 
| 24333 |    569 |   by (clarsimp simp add: word_mod_def uint_2p and_mask_mod_2p) 
 | 
|  |    570 | 
 | 
|  |    571 | lemmas word_mod_2p_is_mask = refl [THEN word_mod_2p_is_mask'] 
 | 
|  |    572 | 
 | 
|  |    573 | lemma mask_eqs:
 | 
|  |    574 |   "(a AND mask n) + b AND mask n = a + b AND mask n"
 | 
|  |    575 |   "a + (b AND mask n) AND mask n = a + b AND mask n"
 | 
|  |    576 |   "(a AND mask n) - b AND mask n = a - b AND mask n"
 | 
|  |    577 |   "a - (b AND mask n) AND mask n = a - b AND mask n"
 | 
|  |    578 |   "a * (b AND mask n) AND mask n = a * b AND mask n"
 | 
|  |    579 |   "(b AND mask n) * a AND mask n = b * a AND mask n"
 | 
|  |    580 |   "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n"
 | 
|  |    581 |   "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n"
 | 
|  |    582 |   "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n"
 | 
|  |    583 |   "- (a AND mask n) AND mask n = - a AND mask n"
 | 
|  |    584 |   "word_succ (a AND mask n) AND mask n = word_succ a AND mask n"
 | 
|  |    585 |   "word_pred (a AND mask n) AND mask n = word_pred a AND mask n"
 | 
|  |    586 |   using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b]
 | 
|  |    587 |   by (auto simp: and_mask_wi bintr_ariths bintr_arith1s new_word_of_int_homs)
 | 
|  |    588 | 
 | 
|  |    589 | lemma mask_power_eq:
 | 
|  |    590 |   "(x AND mask n) ^ k AND mask n = x ^ k AND mask n"
 | 
|  |    591 |   using word_of_int_Ex [where x=x]
 | 
|  |    592 |   by (clarsimp simp: and_mask_wi word_of_int_power_hom bintr_ariths)
 | 
|  |    593 | 
 | 
|  |    594 | 
 | 
| 24350 |    595 | subsubsection "Revcast"
 | 
| 24333 |    596 | 
 | 
|  |    597 | lemmas revcast_def' = revcast_def [simplified]
 | 
|  |    598 | lemmas revcast_def'' = revcast_def' [simplified word_size]
 | 
|  |    599 | lemmas revcast_no_def [simp] =
 | 
|  |    600 |   revcast_def' [where w="number_of ?w", unfolded word_size]
 | 
|  |    601 | 
 | 
|  |    602 | lemma to_bl_revcast: 
 | 
| 24465 |    603 |   "to_bl (revcast w :: 'a :: len0 word) = 
 | 
|  |    604 |    takefill False (len_of TYPE ('a)) (to_bl w)"
 | 
| 24333 |    605 |   apply (unfold revcast_def' word_size)
 | 
|  |    606 |   apply (rule word_bl.Abs_inverse)
 | 
|  |    607 |   apply simp
 | 
|  |    608 |   done
 | 
|  |    609 | 
 | 
|  |    610 | lemma revcast_rev_ucast': 
 | 
|  |    611 |   "cs = [rc, uc] ==> rc = revcast (word_reverse w) ==> uc = ucast w ==> 
 | 
|  |    612 |     rc = word_reverse uc"
 | 
|  |    613 |   apply (unfold ucast_def revcast_def' Let_def word_reverse_def)
 | 
|  |    614 |   apply (clarsimp simp add : to_bl_of_bin takefill_bintrunc)
 | 
|  |    615 |   apply (simp add : word_bl.Abs_inverse word_size)
 | 
|  |    616 |   done
 | 
|  |    617 | 
 | 
|  |    618 | lemmas revcast_rev_ucast = revcast_rev_ucast' [OF refl refl refl]
 | 
|  |    619 | 
 | 
|  |    620 | lemmas revcast_ucast = revcast_rev_ucast
 | 
|  |    621 |   [where w = "word_reverse ?w1", simplified word_rev_rev, standard]
 | 
|  |    622 | 
 | 
|  |    623 | lemmas ucast_revcast = revcast_rev_ucast [THEN word_rev_gal', standard]
 | 
|  |    624 | lemmas ucast_rev_revcast = revcast_ucast [THEN word_rev_gal', standard]
 | 
|  |    625 | 
 | 
|  |    626 | 
 | 
|  |    627 | -- "linking revcast and cast via shift"
 | 
|  |    628 | 
 | 
|  |    629 | lemmas wsst_TYs = source_size target_size word_size
 | 
|  |    630 | 
 | 
|  |    631 | lemma revcast_down_uu': 
 | 
|  |    632 |   "rc = revcast ==> source_size rc = target_size rc + n ==> 
 | 
| 24465 |    633 |     rc (w :: 'a :: len word) = ucast (w >> n)"
 | 
| 24333 |    634 |   apply (simp add: revcast_def')
 | 
|  |    635 |   apply (rule word_bl.Rep_inverse')
 | 
|  |    636 |   apply (rule trans, rule ucast_down_drop)
 | 
|  |    637 |    prefer 2
 | 
|  |    638 |    apply (rule trans, rule drop_shiftr)
 | 
|  |    639 |    apply (auto simp: takefill_alt wsst_TYs)
 | 
|  |    640 |   done
 | 
|  |    641 | 
 | 
|  |    642 | lemma revcast_down_us': 
 | 
|  |    643 |   "rc = revcast ==> source_size rc = target_size rc + n ==> 
 | 
| 24465 |    644 |     rc (w :: 'a :: len word) = ucast (w >>> n)"
 | 
| 24333 |    645 |   apply (simp add: revcast_def')
 | 
|  |    646 |   apply (rule word_bl.Rep_inverse')
 | 
|  |    647 |   apply (rule trans, rule ucast_down_drop)
 | 
|  |    648 |    prefer 2
 | 
|  |    649 |    apply (rule trans, rule drop_sshiftr)
 | 
|  |    650 |    apply (auto simp: takefill_alt wsst_TYs)
 | 
|  |    651 |   done
 | 
|  |    652 | 
 | 
|  |    653 | lemma revcast_down_su': 
 | 
|  |    654 |   "rc = revcast ==> source_size rc = target_size rc + n ==> 
 | 
| 24465 |    655 |     rc (w :: 'a :: len word) = scast (w >> n)"
 | 
| 24333 |    656 |   apply (simp add: revcast_def')
 | 
|  |    657 |   apply (rule word_bl.Rep_inverse')
 | 
|  |    658 |   apply (rule trans, rule scast_down_drop)
 | 
|  |    659 |    prefer 2
 | 
|  |    660 |    apply (rule trans, rule drop_shiftr)
 | 
|  |    661 |    apply (auto simp: takefill_alt wsst_TYs)
 | 
|  |    662 |   done
 | 
|  |    663 | 
 | 
|  |    664 | lemma revcast_down_ss': 
 | 
|  |    665 |   "rc = revcast ==> source_size rc = target_size rc + n ==> 
 | 
| 24465 |    666 |     rc (w :: 'a :: len word) = scast (w >>> n)"
 | 
| 24333 |    667 |   apply (simp add: revcast_def')
 | 
|  |    668 |   apply (rule word_bl.Rep_inverse')
 | 
|  |    669 |   apply (rule trans, rule scast_down_drop)
 | 
|  |    670 |    prefer 2
 | 
|  |    671 |    apply (rule trans, rule drop_sshiftr)
 | 
|  |    672 |    apply (auto simp: takefill_alt wsst_TYs)
 | 
|  |    673 |   done
 | 
|  |    674 | 
 | 
|  |    675 | lemmas revcast_down_uu = refl [THEN revcast_down_uu']
 | 
|  |    676 | lemmas revcast_down_us = refl [THEN revcast_down_us']
 | 
|  |    677 | lemmas revcast_down_su = refl [THEN revcast_down_su']
 | 
|  |    678 | lemmas revcast_down_ss = refl [THEN revcast_down_ss']
 | 
|  |    679 | 
 | 
|  |    680 | lemma cast_down_rev: 
 | 
|  |    681 |   "uc = ucast ==> source_size uc = target_size uc + n ==> 
 | 
| 24465 |    682 |     uc w = revcast ((w :: 'a :: len word) << n)"
 | 
| 24333 |    683 |   apply (unfold shiftl_rev)
 | 
|  |    684 |   apply clarify
 | 
|  |    685 |   apply (simp add: revcast_rev_ucast)
 | 
|  |    686 |   apply (rule word_rev_gal')
 | 
|  |    687 |   apply (rule trans [OF _ revcast_rev_ucast])
 | 
|  |    688 |   apply (rule revcast_down_uu [symmetric])
 | 
|  |    689 |   apply (auto simp add: wsst_TYs)
 | 
|  |    690 |   done
 | 
|  |    691 | 
 | 
|  |    692 | lemma revcast_up': 
 | 
|  |    693 |   "rc = revcast ==> source_size rc + n = target_size rc ==> 
 | 
| 24465 |    694 |     rc w = (ucast w :: 'a :: len word) << n" 
 | 
| 24333 |    695 |   apply (simp add: revcast_def')
 | 
|  |    696 |   apply (rule word_bl.Rep_inverse')
 | 
|  |    697 |   apply (simp add: takefill_alt)
 | 
|  |    698 |   apply (rule bl_shiftl [THEN trans])
 | 
|  |    699 |   apply (subst ucast_up_app)
 | 
|  |    700 |   apply (auto simp add: wsst_TYs)
 | 
|  |    701 |   apply (drule sym)
 | 
|  |    702 |   apply (simp add: min_def)
 | 
|  |    703 |   done
 | 
|  |    704 | 
 | 
|  |    705 | lemmas revcast_up = refl [THEN revcast_up']
 | 
|  |    706 | 
 | 
|  |    707 | lemmas rc1 = revcast_up [THEN 
 | 
|  |    708 |   revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]]
 | 
|  |    709 | lemmas rc2 = revcast_down_uu [THEN 
 | 
|  |    710 |   revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]]
 | 
|  |    711 | 
 | 
|  |    712 | lemmas ucast_up =
 | 
|  |    713 |  rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]]
 | 
|  |    714 | lemmas ucast_down = 
 | 
|  |    715 |   rc2 [simplified rev_shiftr revcast_ucast [symmetric]]
 | 
|  |    716 | 
 | 
|  |    717 | 
 | 
| 24350 |    718 | subsubsection "Slices"
 | 
| 24333 |    719 | 
 | 
|  |    720 | lemmas slice1_no_bin [simp] =
 | 
|  |    721 |   slice1_def [where w="number_of ?w", unfolded to_bl_no_bin]
 | 
|  |    722 | 
 | 
|  |    723 | lemmas slice_no_bin [simp] = 
 | 
|  |    724 |    trans [OF slice_def [THEN meta_eq_to_obj_eq] 
 | 
|  |    725 |              slice1_no_bin [THEN meta_eq_to_obj_eq], 
 | 
|  |    726 |           unfolded word_size, standard]
 | 
|  |    727 | 
 | 
|  |    728 | lemma slice1_0 [simp] : "slice1 n 0 = 0"
 | 
|  |    729 |   unfolding slice1_def by (simp add : to_bl_0)
 | 
|  |    730 | 
 | 
|  |    731 | lemma slice_0 [simp] : "slice n 0 = 0"
 | 
|  |    732 |   unfolding slice_def by auto
 | 
|  |    733 | 
 | 
|  |    734 | lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))"
 | 
|  |    735 |   unfolding slice_def' slice1_def
 | 
|  |    736 |   by (simp add : takefill_alt word_size)
 | 
|  |    737 | 
 | 
|  |    738 | lemmas slice_take = slice_take' [unfolded word_size]
 | 
|  |    739 | 
 | 
|  |    740 | -- "shiftr to a word of the same size is just slice, 
 | 
|  |    741 |     slice is just shiftr then ucast"
 | 
|  |    742 | lemmas shiftr_slice = trans
 | 
|  |    743 |   [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric], standard]
 | 
|  |    744 | 
 | 
|  |    745 | lemma slice_shiftr: "slice n w = ucast (w >> n)"
 | 
|  |    746 |   apply (unfold slice_take shiftr_bl)
 | 
|  |    747 |   apply (rule ucast_of_bl_up [symmetric])
 | 
|  |    748 |   apply (simp add: word_size)
 | 
|  |    749 |   done
 | 
|  |    750 | 
 | 
|  |    751 | lemma nth_slice: 
 | 
| 24465 |    752 |   "(slice n w :: 'a :: len0 word) !! m = 
 | 
|  |    753 |    (w !! (m + n) & m < len_of TYPE ('a))"
 | 
| 24333 |    754 |   unfolding slice_shiftr 
 | 
|  |    755 |   by (simp add : nth_ucast nth_shiftr)
 | 
|  |    756 | 
 | 
|  |    757 | lemma slice1_down_alt': 
 | 
|  |    758 |   "sl = slice1 n w ==> fs = size sl ==> fs + k = n ==> 
 | 
|  |    759 |     to_bl sl = takefill False fs (drop k (to_bl w))"
 | 
|  |    760 |   unfolding slice1_def word_size of_bl_def uint_bl
 | 
|  |    761 |   by (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop drop_takefill)
 | 
|  |    762 | 
 | 
|  |    763 | lemma slice1_up_alt': 
 | 
|  |    764 |   "sl = slice1 n w ==> fs = size sl ==> fs = n + k ==> 
 | 
|  |    765 |     to_bl sl = takefill False fs (replicate k False @ (to_bl w))"
 | 
|  |    766 |   apply (unfold slice1_def word_size of_bl_def uint_bl)
 | 
|  |    767 |   apply (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop 
 | 
|  |    768 |                         takefill_append [symmetric])
 | 
| 24465 |    769 |   apply (rule_tac f = "%k. takefill False (len_of TYPE('a))
 | 
|  |    770 |     (replicate k False @ bin_to_bl (len_of TYPE('b)) (uint w))" in arg_cong)
 | 
| 24333 |    771 |   apply arith
 | 
|  |    772 |   done
 | 
|  |    773 |     
 | 
|  |    774 | lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size]
 | 
|  |    775 | lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size]
 | 
|  |    776 | lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1]
 | 
|  |    777 | lemmas slice1_up_alts = 
 | 
|  |    778 |   le_add_diff_inverse [symmetric, THEN su1] 
 | 
|  |    779 |   le_add_diff_inverse2 [symmetric, THEN su1]
 | 
|  |    780 | 
 | 
|  |    781 | lemma ucast_slice1: "ucast w = slice1 (size w) w"
 | 
|  |    782 |   unfolding slice1_def ucast_bl
 | 
|  |    783 |   by (simp add : takefill_same' word_size)
 | 
|  |    784 | 
 | 
|  |    785 | lemma ucast_slice: "ucast w = slice 0 w"
 | 
|  |    786 |   unfolding slice_def by (simp add : ucast_slice1)
 | 
|  |    787 | 
 | 
|  |    788 | lemmas slice_id = trans [OF ucast_slice [symmetric] ucast_id]
 | 
|  |    789 | 
 | 
|  |    790 | lemma revcast_slice1': 
 | 
|  |    791 |   "rc = revcast w ==> slice1 (size rc) w = rc"
 | 
|  |    792 |   unfolding slice1_def revcast_def' by (simp add : word_size)
 | 
|  |    793 | 
 | 
|  |    794 | lemmas revcast_slice1 = refl [THEN revcast_slice1']
 | 
|  |    795 | 
 | 
|  |    796 | lemma slice1_tf_tf': 
 | 
| 24465 |    797 |   "to_bl (slice1 n w :: 'a :: len0 word) = 
 | 
|  |    798 |    rev (takefill False (len_of TYPE('a)) (rev (takefill False n (to_bl w))))"
 | 
| 24333 |    799 |   unfolding slice1_def by (rule word_rev_tf)
 | 
|  |    800 | 
 | 
|  |    801 | lemmas slice1_tf_tf = slice1_tf_tf'
 | 
|  |    802 |   [THEN word_bl.Rep_inverse', symmetric, standard]
 | 
|  |    803 | 
 | 
|  |    804 | lemma rev_slice1:
 | 
| 24465 |    805 |   "n + k = len_of TYPE('a) + len_of TYPE('b) \<Longrightarrow> 
 | 
|  |    806 |   slice1 n (word_reverse w :: 'b :: len0 word) = 
 | 
|  |    807 |   word_reverse (slice1 k w :: 'a :: len0 word)"
 | 
| 24333 |    808 |   apply (unfold word_reverse_def slice1_tf_tf)
 | 
|  |    809 |   apply (rule word_bl.Rep_inverse')
 | 
|  |    810 |   apply (rule rev_swap [THEN iffD1])
 | 
|  |    811 |   apply (rule trans [symmetric])
 | 
|  |    812 |   apply (rule tf_rev)
 | 
|  |    813 |    apply (simp add: word_bl.Abs_inverse)
 | 
|  |    814 |   apply (simp add: word_bl.Abs_inverse)
 | 
|  |    815 |   done
 | 
|  |    816 | 
 | 
|  |    817 | lemma rev_slice': 
 | 
|  |    818 |   "res = slice n (word_reverse w) ==> n + k + size res = size w ==> 
 | 
|  |    819 |     res = word_reverse (slice k w)"
 | 
|  |    820 |   apply (unfold slice_def word_size)
 | 
|  |    821 |   apply clarify
 | 
|  |    822 |   apply (rule rev_slice1)
 | 
|  |    823 |   apply arith
 | 
|  |    824 |   done
 | 
|  |    825 | 
 | 
|  |    826 | lemmas rev_slice = refl [THEN rev_slice', unfolded word_size]
 | 
|  |    827 | 
 | 
|  |    828 | lemmas sym_notr = 
 | 
|  |    829 |   not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]]
 | 
|  |    830 | 
 | 
|  |    831 | -- {* problem posed by TPHOLs referee:
 | 
|  |    832 |       criterion for overflow of addition of signed integers *}
 | 
|  |    833 | 
 | 
|  |    834 | lemma sofl_test:
 | 
| 24465 |    835 |   "(sint (x :: 'a :: len word) + sint y = sint (x + y)) = 
 | 
| 24333 |    836 |      ((((x+y) XOR x) AND ((x+y) XOR y)) >> (size x - 1) = 0)"
 | 
|  |    837 |   apply (unfold word_size)
 | 
| 24465 |    838 |   apply (cases "len_of TYPE('a)", simp) 
 | 
| 24333 |    839 |   apply (subst msb_shift [THEN sym_notr])
 | 
|  |    840 |   apply (simp add: word_ops_msb)
 | 
|  |    841 |   apply (simp add: word_msb_sint)
 | 
|  |    842 |   apply safe
 | 
|  |    843 |        apply simp_all
 | 
|  |    844 |      apply (unfold sint_word_ariths)
 | 
|  |    845 |      apply (unfold word_sbin.set_iff_norm [symmetric] sints_num)
 | 
|  |    846 |      apply safe
 | 
|  |    847 |         apply (insert sint_range' [where x=x])
 | 
|  |    848 |         apply (insert sint_range' [where x=y])
 | 
|  |    849 |         defer 
 | 
|  |    850 |         apply (simp (no_asm), arith)
 | 
|  |    851 |        apply (simp (no_asm), arith)
 | 
|  |    852 |       defer
 | 
|  |    853 |       defer
 | 
|  |    854 |       apply (simp (no_asm), arith)
 | 
|  |    855 |      apply (simp (no_asm), arith)
 | 
|  |    856 |     apply (rule notI [THEN notnotD],
 | 
|  |    857 |            drule leI not_leE,
 | 
|  |    858 |            drule sbintrunc_inc sbintrunc_dec,      
 | 
|  |    859 |            simp)+
 | 
|  |    860 |   done
 | 
|  |    861 | 
 | 
|  |    862 | 
 | 
| 24350 |    863 | subsection "Split and cat"
 | 
| 24333 |    864 | 
 | 
|  |    865 | lemmas word_split_bin' = word_split_def [THEN meta_eq_to_obj_eq, standard]
 | 
|  |    866 | lemmas word_cat_bin' = word_cat_def [THEN meta_eq_to_obj_eq, standard]
 | 
|  |    867 | 
 | 
|  |    868 | lemma word_rsplit_no:
 | 
| 24465 |    869 |   "(word_rsplit (number_of bin :: 'b :: len0 word) :: 'a word list) = 
 | 
|  |    870 |     map number_of (bin_rsplit (len_of TYPE('a :: len)) 
 | 
|  |    871 |       (len_of TYPE('b), bintrunc (len_of TYPE('b)) bin))"
 | 
| 24333 |    872 |   apply (unfold word_rsplit_def word_no_wi)
 | 
|  |    873 |   apply (simp add: word_ubin.eq_norm)
 | 
|  |    874 |   done
 | 
|  |    875 | 
 | 
|  |    876 | lemmas word_rsplit_no_cl [simp] = word_rsplit_no
 | 
|  |    877 |   [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]]
 | 
|  |    878 | 
 | 
|  |    879 | lemma test_bit_cat:
 | 
|  |    880 |   "wc = word_cat a b ==> wc !! n = (n < size wc & 
 | 
|  |    881 |     (if n < size b then b !! n else a !! (n - size b)))"
 | 
|  |    882 |   apply (unfold word_cat_bin' test_bit_bin)
 | 
|  |    883 |   apply (auto simp add : word_ubin.eq_norm nth_bintr bin_nth_cat word_size)
 | 
|  |    884 |   apply (erule bin_nth_uint_imp)
 | 
|  |    885 |   done
 | 
|  |    886 | 
 | 
|  |    887 | lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)"
 | 
|  |    888 |   apply (unfold of_bl_def to_bl_def word_cat_bin')
 | 
|  |    889 |   apply (simp add: bl_to_bin_app_cat)
 | 
|  |    890 |   done
 | 
|  |    891 | 
 | 
|  |    892 | lemma of_bl_append:
 | 
| 24465 |    893 |   "(of_bl (xs @ ys) :: 'a :: len word) = of_bl xs * 2^(length ys) + of_bl ys"
 | 
| 24333 |    894 |   apply (unfold of_bl_def)
 | 
|  |    895 |   apply (simp add: bl_to_bin_app_cat bin_cat_num)
 | 
|  |    896 |   apply (simp add: word_of_int_power_hom [symmetric] new_word_of_int_hom_syms)
 | 
|  |    897 |   done
 | 
|  |    898 | 
 | 
|  |    899 | lemma of_bl_False [simp]:
 | 
|  |    900 |   "of_bl (False#xs) = of_bl xs"
 | 
|  |    901 |   by (rule word_eqI)
 | 
|  |    902 |      (auto simp add: test_bit_of_bl nth_append)
 | 
|  |    903 | 
 | 
|  |    904 | lemma of_bl_True: 
 | 
| 24465 |    905 |   "(of_bl (True#xs)::'a::len word) = 2^length xs + of_bl xs"
 | 
| 24333 |    906 |   by (subst of_bl_append [where xs="[True]", simplified])
 | 
|  |    907 |      (simp add: word_1_bl)
 | 
|  |    908 | 
 | 
|  |    909 | lemma of_bl_Cons:
 | 
|  |    910 |   "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs"
 | 
|  |    911 |   by (cases x) (simp_all add: of_bl_True)
 | 
|  |    912 | 
 | 
| 24465 |    913 | lemma split_uint_lem: "bin_split n (uint (w :: 'a :: len0 word)) = (a, b) ==> 
 | 
|  |    914 |   a = bintrunc (len_of TYPE('a) - n) a & b = bintrunc (len_of TYPE('a)) b"
 | 
| 24333 |    915 |   apply (frule word_ubin.norm_Rep [THEN ssubst])
 | 
|  |    916 |   apply (drule bin_split_trunc1)
 | 
|  |    917 |   apply (drule sym [THEN trans])
 | 
|  |    918 |   apply assumption
 | 
|  |    919 |   apply safe
 | 
|  |    920 |   done
 | 
|  |    921 | 
 | 
|  |    922 | lemma word_split_bl': 
 | 
|  |    923 |   "std = size c - size b ==> (word_split c = (a, b)) ==> 
 | 
|  |    924 |     (a = of_bl (take std (to_bl c)) & b = of_bl (drop std (to_bl c)))"
 | 
|  |    925 |   apply (unfold word_split_bin')
 | 
|  |    926 |   apply safe
 | 
|  |    927 |    defer
 | 
|  |    928 |    apply (clarsimp split: prod.splits)
 | 
|  |    929 |    apply (drule word_ubin.norm_Rep [THEN ssubst])
 | 
|  |    930 |    apply (drule split_bintrunc)
 | 
|  |    931 |    apply (simp add : of_bl_def bl2bin_drop word_size
 | 
|  |    932 |        word_ubin.norm_eq_iff [symmetric] min_def del : word_ubin.norm_Rep)    
 | 
|  |    933 |   apply (clarsimp split: prod.splits)
 | 
|  |    934 |   apply (frule split_uint_lem [THEN conjunct1])
 | 
|  |    935 |   apply (unfold word_size)
 | 
| 24465 |    936 |   apply (cases "len_of TYPE('a) >= len_of TYPE('b)")
 | 
| 24333 |    937 |    defer
 | 
|  |    938 |    apply (simp add: word_0_bl word_0_wi_Pls)
 | 
|  |    939 |   apply (simp add : of_bl_def to_bl_def)
 | 
|  |    940 |   apply (subst bin_split_take1 [symmetric])
 | 
|  |    941 |     prefer 2
 | 
|  |    942 |     apply assumption
 | 
|  |    943 |    apply simp
 | 
|  |    944 |   apply (erule thin_rl)
 | 
|  |    945 |   apply (erule arg_cong [THEN trans])
 | 
|  |    946 |   apply (simp add : word_ubin.norm_eq_iff [symmetric])
 | 
|  |    947 |   done
 | 
|  |    948 | 
 | 
|  |    949 | lemma word_split_bl: "std = size c - size b ==> 
 | 
|  |    950 |     (a = of_bl (take std (to_bl c)) & b = of_bl (drop std (to_bl c))) <-> 
 | 
|  |    951 |     word_split c = (a, b)"
 | 
|  |    952 |   apply (rule iffI)
 | 
|  |    953 |    defer
 | 
|  |    954 |    apply (erule (1) word_split_bl')
 | 
|  |    955 |   apply (case_tac "word_split c")
 | 
|  |    956 |   apply (auto simp add : word_size)
 | 
|  |    957 |   apply (frule word_split_bl' [rotated])
 | 
|  |    958 |   apply (auto simp add : word_size)
 | 
|  |    959 |   done
 | 
|  |    960 | 
 | 
|  |    961 | lemma word_split_bl_eq:
 | 
| 24465 |    962 |    "(word_split (c::'a::len word) :: ('c :: len0 word * 'd :: len0 word)) =
 | 
|  |    963 |       (of_bl (take (len_of TYPE('a::len) - len_of TYPE('d::len0)) (to_bl c)),
 | 
|  |    964 |        of_bl (drop (len_of TYPE('a) - len_of TYPE('d)) (to_bl c)))"
 | 
| 24333 |    965 |   apply (rule word_split_bl [THEN iffD1])
 | 
|  |    966 |   apply (unfold word_size)
 | 
|  |    967 |   apply (rule refl conjI)+
 | 
|  |    968 |   done
 | 
|  |    969 | 
 | 
|  |    970 | -- "keep quantifiers for use in simplification"
 | 
|  |    971 | lemma test_bit_split':
 | 
|  |    972 |   "word_split c = (a, b) --> (ALL n m. b !! n = (n < size b & c !! n) & 
 | 
|  |    973 |     a !! m = (m < size a & c !! (m + size b)))"
 | 
|  |    974 |   apply (unfold word_split_bin' test_bit_bin)
 | 
|  |    975 |   apply (clarify)
 | 
|  |    976 |   apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits)
 | 
|  |    977 |   apply (drule bin_nth_split)
 | 
|  |    978 |   apply safe
 | 
|  |    979 |        apply (simp_all add: add_commute)
 | 
|  |    980 |    apply (erule bin_nth_uint_imp)+
 | 
|  |    981 |   done
 | 
|  |    982 | 
 | 
|  |    983 | lemmas test_bit_split = 
 | 
|  |    984 |   test_bit_split' [THEN mp, simplified all_simps, standard]
 | 
|  |    985 | 
 | 
|  |    986 | lemma test_bit_split_eq: "word_split c = (a, b) <-> 
 | 
|  |    987 |   ((ALL n::nat. b !! n = (n < size b & c !! n)) &
 | 
|  |    988 |     (ALL m::nat. a !! m = (m < size a & c !! (m + size b))))"
 | 
|  |    989 |   apply (rule_tac iffI)
 | 
|  |    990 |    apply (rule_tac conjI)
 | 
|  |    991 |     apply (erule test_bit_split [THEN conjunct1])
 | 
|  |    992 |    apply (erule test_bit_split [THEN conjunct2])
 | 
|  |    993 |   apply (case_tac "word_split c")
 | 
|  |    994 |   apply (frule test_bit_split)
 | 
|  |    995 |   apply (erule trans)
 | 
|  |    996 |   apply (fastsimp intro ! : word_eqI simp add : word_size)
 | 
|  |    997 |   done
 | 
|  |    998 | 
 | 
|  |    999 | -- {* this odd result is analogous to ucast\_id, 
 | 
|  |   1000 |       result to the length given by the result type *}
 | 
|  |   1001 | 
 | 
|  |   1002 | lemma word_cat_id: "word_cat a b = b"
 | 
|  |   1003 |   unfolding word_cat_bin' by (simp add: word_ubin.inverse_norm)
 | 
|  |   1004 | 
 | 
|  |   1005 | -- "limited hom result"
 | 
|  |   1006 | lemma word_cat_hom:
 | 
| 24465 |   1007 |   "len_of TYPE('a::len0) <= len_of TYPE('b::len0) + len_of TYPE ('c::len0)
 | 
| 24333 |   1008 |   ==>
 | 
|  |   1009 |   (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = 
 | 
|  |   1010 |   word_of_int (bin_cat w (size b) (uint b))"
 | 
|  |   1011 |   apply (unfold word_cat_def word_size) 
 | 
|  |   1012 |   apply (clarsimp simp add : word_ubin.norm_eq_iff [symmetric]
 | 
|  |   1013 |       word_ubin.eq_norm bintr_cat min_def)
 | 
| 24465 |   1014 |   apply arith
 | 
| 24333 |   1015 |   done
 | 
|  |   1016 | 
 | 
|  |   1017 | lemma word_cat_split_alt:
 | 
|  |   1018 |   "size w <= size u + size v ==> word_split w = (u, v) ==> word_cat u v = w"
 | 
|  |   1019 |   apply (rule word_eqI)
 | 
|  |   1020 |   apply (drule test_bit_split)
 | 
|  |   1021 |   apply (clarsimp simp add : test_bit_cat word_size)
 | 
|  |   1022 |   apply safe
 | 
|  |   1023 |   apply arith
 | 
|  |   1024 |   done
 | 
|  |   1025 | 
 | 
|  |   1026 | lemmas word_cat_split_size = 
 | 
|  |   1027 |   sym [THEN [2] word_cat_split_alt [symmetric], standard]
 | 
|  |   1028 | 
 | 
|  |   1029 | 
 | 
| 24350 |   1030 | subsubsection "Split and slice"
 | 
| 24333 |   1031 | 
 | 
|  |   1032 | lemma split_slices: 
 | 
|  |   1033 |   "word_split w = (u, v) ==> u = slice (size v) w & v = slice 0 w"
 | 
|  |   1034 |   apply (drule test_bit_split)
 | 
|  |   1035 |   apply (rule conjI)
 | 
|  |   1036 |    apply (rule word_eqI, clarsimp simp: nth_slice word_size)+
 | 
|  |   1037 |   done
 | 
|  |   1038 | 
 | 
|  |   1039 | lemma slice_cat1':
 | 
|  |   1040 |   "wc = word_cat a b ==> size wc >= size a + size b ==> slice (size b) wc = a"
 | 
|  |   1041 |   apply safe
 | 
|  |   1042 |   apply (rule word_eqI)
 | 
|  |   1043 |   apply (simp add: nth_slice test_bit_cat word_size)
 | 
|  |   1044 |   done
 | 
|  |   1045 | 
 | 
|  |   1046 | lemmas slice_cat1 = refl [THEN slice_cat1']
 | 
|  |   1047 | lemmas slice_cat2 = trans [OF slice_id word_cat_id]
 | 
|  |   1048 | 
 | 
|  |   1049 | lemma cat_slices:
 | 
|  |   1050 |   "a = slice n c ==> b = slice 0 c ==> n = size b ==> \  
 | 
|  |   1051 |     size a + size b >= size c ==> word_cat a b = c"
 | 
|  |   1052 |   apply safe
 | 
|  |   1053 |   apply (rule word_eqI)
 | 
|  |   1054 |   apply (simp add: nth_slice test_bit_cat word_size)
 | 
|  |   1055 |   apply safe
 | 
|  |   1056 |   apply arith
 | 
|  |   1057 |   done
 | 
|  |   1058 | 
 | 
|  |   1059 | lemma word_split_cat_alt:
 | 
|  |   1060 |   "w = word_cat u v ==> size u + size v <= size w ==> word_split w = (u, v)"
 | 
|  |   1061 |   apply (case_tac "word_split ?w")
 | 
|  |   1062 |   apply (rule trans, assumption)
 | 
|  |   1063 |   apply (drule test_bit_split)
 | 
|  |   1064 |   apply safe
 | 
|  |   1065 |    apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+
 | 
|  |   1066 |   done
 | 
|  |   1067 | 
 | 
|  |   1068 | lemmas word_cat_bl_no_bin [simp] =
 | 
|  |   1069 |   word_cat_bl [where a="number_of ?a" 
 | 
|  |   1070 |                  and b="number_of ?b", 
 | 
|  |   1071 |                unfolded to_bl_no_bin]
 | 
|  |   1072 | 
 | 
|  |   1073 | lemmas word_split_bl_no_bin [simp] =
 | 
|  |   1074 |   word_split_bl_eq [where c="number_of ?c", unfolded to_bl_no_bin]
 | 
|  |   1075 | 
 | 
|  |   1076 | -- {* this odd result arises from the fact that the statement of the
 | 
|  |   1077 |       result implies that the decoded words are of the same type, 
 | 
|  |   1078 |       and therefore of the same length, as the original word *}
 | 
|  |   1079 | 
 | 
|  |   1080 | lemma word_rsplit_same: "word_rsplit w = [w]"
 | 
|  |   1081 |   unfolding word_rsplit_def by (simp add : bin_rsplit_all)
 | 
|  |   1082 | 
 | 
|  |   1083 | lemma word_rsplit_empty_iff_size:
 | 
|  |   1084 |   "(word_rsplit w = []) = (size w = 0)" 
 | 
|  |   1085 |   unfolding word_rsplit_def bin_rsplit_def word_size
 | 
|  |   1086 |   by (simp add: bin_rsplit_aux_simp_alt Let_def split: split_split)
 | 
|  |   1087 | 
 | 
|  |   1088 | lemma test_bit_rsplit:
 | 
| 24465 |   1089 |   "sw = word_rsplit w ==> m < size (hd sw :: 'a :: len word) ==> 
 | 
| 24333 |   1090 |     k < length sw ==> (rev sw ! k) !! m = (w !! (k * size (hd sw) + m))"
 | 
|  |   1091 |   apply (unfold word_rsplit_def word_test_bit_def)
 | 
|  |   1092 |   apply (rule trans)
 | 
|  |   1093 |    apply (rule_tac f = "%x. bin_nth x m" in arg_cong)
 | 
|  |   1094 |    apply (rule nth_map [symmetric])
 | 
|  |   1095 |    apply simp
 | 
|  |   1096 |   apply (rule bin_nth_rsplit)
 | 
|  |   1097 |      apply simp_all
 | 
|  |   1098 |   apply (simp add : word_size rev_map map_compose [symmetric])
 | 
|  |   1099 |   apply (rule trans)
 | 
|  |   1100 |    defer
 | 
|  |   1101 |    apply (rule map_ident [THEN fun_cong])
 | 
|  |   1102 |   apply (rule refl [THEN map_cong])
 | 
|  |   1103 |   apply (simp add : word_ubin.eq_norm)
 | 
| 24465 |   1104 |   apply (erule bin_rsplit_size_sign [OF len_gt_0 refl])
 | 
| 24333 |   1105 |   done
 | 
|  |   1106 | 
 | 
|  |   1107 | lemma word_rcat_bl: "word_rcat wl == of_bl (concat (map to_bl wl))"
 | 
|  |   1108 |   unfolding word_rcat_def to_bl_def' of_bl_def
 | 
|  |   1109 |   by (clarsimp simp add : bin_rcat_bl map_compose)
 | 
|  |   1110 | 
 | 
|  |   1111 | lemma size_rcat_lem':
 | 
|  |   1112 |   "size (concat (map to_bl wl)) = length wl * size (hd wl)"
 | 
|  |   1113 |   unfolding word_size by (induct wl) auto
 | 
|  |   1114 | 
 | 
|  |   1115 | lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size]
 | 
|  |   1116 | 
 | 
| 24465 |   1117 | lemmas td_gal_lt_len = len_gt_0 [THEN td_gal_lt, standard]
 | 
| 24333 |   1118 | 
 | 
|  |   1119 | lemma nth_rcat_lem' [rule_format] :
 | 
| 24465 |   1120 |   "sw = size (hd wl  :: 'a :: len word) ==> (ALL n. n < size wl * sw --> 
 | 
| 24333 |   1121 |     rev (concat (map to_bl wl)) ! n = 
 | 
|  |   1122 |     rev (to_bl (rev wl ! (n div sw))) ! (n mod sw))"
 | 
|  |   1123 |   apply (unfold word_size)
 | 
|  |   1124 |   apply (induct "wl")
 | 
|  |   1125 |    apply clarsimp
 | 
|  |   1126 |   apply (clarsimp simp add : nth_append size_rcat_lem)
 | 
|  |   1127 |   apply (simp (no_asm_use) only:  mult_Suc [symmetric] 
 | 
|  |   1128 |          td_gal_lt_len less_Suc_eq_le mod_div_equality')
 | 
|  |   1129 |   apply clarsimp
 | 
|  |   1130 |   done
 | 
|  |   1131 | 
 | 
|  |   1132 | lemmas nth_rcat_lem = refl [THEN nth_rcat_lem', unfolded word_size]
 | 
|  |   1133 | 
 | 
|  |   1134 | lemma test_bit_rcat:
 | 
| 24465 |   1135 |   "sw = size (hd wl :: 'a :: len word) ==> rc = word_rcat wl ==> rc !! n = 
 | 
| 24333 |   1136 |     (n < size rc & n div sw < size wl & (rev wl) ! (n div sw) !! (n mod sw))"
 | 
|  |   1137 |   apply (unfold word_rcat_bl word_size)
 | 
|  |   1138 |   apply (clarsimp simp add : 
 | 
|  |   1139 |     test_bit_of_bl size_rcat_lem word_size td_gal_lt_len)
 | 
|  |   1140 |   apply safe
 | 
|  |   1141 |    apply (auto simp add : 
 | 
|  |   1142 |     test_bit_bl word_size td_gal_lt_len [THEN iffD2, THEN nth_rcat_lem])
 | 
|  |   1143 |   done
 | 
|  |   1144 | 
 | 
|  |   1145 | lemma foldl_eq_foldr [rule_format] :
 | 
|  |   1146 |   "ALL x. foldl op + x xs = foldr op + (x # xs) (0 :: 'a :: comm_monoid_add)" 
 | 
|  |   1147 |   by (induct xs) (auto simp add : add_assoc)
 | 
|  |   1148 | 
 | 
|  |   1149 | lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong]
 | 
|  |   1150 | 
 | 
|  |   1151 | lemmas test_bit_rsplit_alt = 
 | 
|  |   1152 |   trans [OF nth_rev_alt [THEN test_bit_cong] 
 | 
|  |   1153 |   test_bit_rsplit [OF refl asm_rl diff_Suc_less]]
 | 
|  |   1154 | 
 | 
|  |   1155 | -- "lazy way of expressing that u and v, and su and sv, have same types"
 | 
|  |   1156 | lemma word_rsplit_len_indep':
 | 
|  |   1157 |   "[u,v] = p ==> [su,sv] = q ==> word_rsplit u = su ==> 
 | 
|  |   1158 |     word_rsplit v = sv ==> length su = length sv"
 | 
|  |   1159 |   apply (unfold word_rsplit_def)
 | 
|  |   1160 |   apply (auto simp add : bin_rsplit_len_indep)
 | 
|  |   1161 |   done
 | 
|  |   1162 | 
 | 
|  |   1163 | lemmas word_rsplit_len_indep = word_rsplit_len_indep' [OF refl refl refl refl]
 | 
|  |   1164 | 
 | 
|  |   1165 | lemma length_word_rsplit_size: 
 | 
| 24465 |   1166 |   "n = len_of TYPE ('a :: len) ==> 
 | 
| 24333 |   1167 |     (length (word_rsplit w :: 'a word list) <= m) = (size w <= m * n)"
 | 
|  |   1168 |   apply (unfold word_rsplit_def word_size)
 | 
|  |   1169 |   apply (clarsimp simp add : bin_rsplit_len_le)
 | 
|  |   1170 |   done
 | 
|  |   1171 | 
 | 
|  |   1172 | lemmas length_word_rsplit_lt_size = 
 | 
|  |   1173 |   length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]]
 | 
|  |   1174 | 
 | 
|  |   1175 | lemma length_word_rsplit_exp_size: 
 | 
| 24465 |   1176 |   "n = len_of TYPE ('a :: len) ==> 
 | 
| 24333 |   1177 |     length (word_rsplit w :: 'a word list) = (size w + n - 1) div n"
 | 
|  |   1178 |   unfolding word_rsplit_def by (clarsimp simp add : word_size bin_rsplit_len)
 | 
|  |   1179 | 
 | 
|  |   1180 | lemma length_word_rsplit_even_size: 
 | 
| 24465 |   1181 |   "n = len_of TYPE ('a :: len) ==> size w = m * n ==> 
 | 
| 24333 |   1182 |     length (word_rsplit w :: 'a word list) = m"
 | 
|  |   1183 |   by (clarsimp simp add : length_word_rsplit_exp_size given_quot_alt)
 | 
|  |   1184 | 
 | 
|  |   1185 | lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size]
 | 
|  |   1186 | 
 | 
|  |   1187 | (* alternative proof of word_rcat_rsplit *)
 | 
|  |   1188 | lemmas tdle = iffD2 [OF split_div_lemma refl, THEN conjunct1] 
 | 
|  |   1189 | lemmas dtle = xtr4 [OF tdle mult_commute]
 | 
|  |   1190 | 
 | 
|  |   1191 | lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w"
 | 
|  |   1192 |   apply (rule word_eqI)
 | 
|  |   1193 |   apply (clarsimp simp add : test_bit_rcat word_size)
 | 
|  |   1194 |   apply (subst refl [THEN test_bit_rsplit])
 | 
|  |   1195 |     apply (simp_all add: word_size 
 | 
|  |   1196 |       refl [THEN length_word_rsplit_size [simplified le_def, simplified]])
 | 
|  |   1197 |    apply safe
 | 
| 24465 |   1198 |    apply (erule xtr7, rule len_gt_0 [THEN dtle])+
 | 
| 24333 |   1199 |   done
 | 
|  |   1200 | 
 | 
|  |   1201 | lemma size_word_rsplit_rcat_size':
 | 
| 24465 |   1202 |   "word_rcat (ws :: 'a :: len word list) = frcw ==> 
 | 
|  |   1203 |     size frcw = length ws * len_of TYPE ('a) ==> 
 | 
| 24333 |   1204 |     size (hd [word_rsplit frcw, ws]) = size ws" 
 | 
|  |   1205 |   apply (clarsimp simp add : word_size length_word_rsplit_exp_size')
 | 
| 24465 |   1206 |   apply (fast intro: given_quot_alt)
 | 
| 24333 |   1207 |   done
 | 
|  |   1208 | 
 | 
|  |   1209 | lemmas size_word_rsplit_rcat_size = 
 | 
|  |   1210 |   size_word_rsplit_rcat_size' [simplified]
 | 
|  |   1211 | 
 | 
|  |   1212 | lemma msrevs:
 | 
|  |   1213 |   fixes n::nat
 | 
|  |   1214 |   shows "0 < n \<Longrightarrow> (k * n + m) div n = m div n + k"
 | 
|  |   1215 |   and   "(k * n + m) mod n = m mod n"
 | 
|  |   1216 |   by (auto simp: add_commute)
 | 
|  |   1217 | 
 | 
|  |   1218 | lemma word_rsplit_rcat_size':
 | 
| 24465 |   1219 |   "word_rcat (ws :: 'a :: len word list) = frcw ==> 
 | 
|  |   1220 |     size frcw = length ws * len_of TYPE ('a) ==> word_rsplit frcw = ws" 
 | 
| 24333 |   1221 |   apply (frule size_word_rsplit_rcat_size, assumption)
 | 
|  |   1222 |   apply (clarsimp simp add : word_size)
 | 
|  |   1223 |   apply (rule nth_equalityI, assumption)
 | 
|  |   1224 |   apply clarsimp
 | 
|  |   1225 |   apply (rule word_eqI)
 | 
|  |   1226 |   apply (rule trans)
 | 
|  |   1227 |    apply (rule test_bit_rsplit_alt)
 | 
|  |   1228 |      apply (clarsimp simp: word_size)+
 | 
|  |   1229 |   apply (rule trans)
 | 
|  |   1230 |   apply (rule test_bit_rcat [OF refl refl])
 | 
|  |   1231 |   apply (simp add : word_size msrevs)
 | 
|  |   1232 |   apply (subst nth_rev)
 | 
|  |   1233 |    apply arith
 | 
|  |   1234 |   apply (simp add : le0 [THEN [2] xtr7, THEN diff_Suc_less])
 | 
|  |   1235 |   apply safe
 | 
|  |   1236 |   apply (simp add : diff_mult_distrib)
 | 
|  |   1237 |   apply (rule mpl_lem)
 | 
|  |   1238 |   apply (cases "size ws")
 | 
|  |   1239 |    apply simp_all
 | 
|  |   1240 |   done
 | 
|  |   1241 | 
 | 
|  |   1242 | lemmas word_rsplit_rcat_size = refl [THEN word_rsplit_rcat_size']
 | 
|  |   1243 | 
 | 
|  |   1244 | 
 | 
| 24350 |   1245 | subsection "Rotation"
 | 
| 24333 |   1246 | 
 | 
|  |   1247 | lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified]
 | 
|  |   1248 | 
 | 
|  |   1249 | lemmas word_rot_defs = word_roti_def word_rotr_def word_rotl_def
 | 
|  |   1250 | 
 | 
|  |   1251 | lemma rotate_eq_mod: 
 | 
|  |   1252 |   "m mod length xs = n mod length xs ==> rotate m xs = rotate n xs"
 | 
|  |   1253 |   apply (rule box_equals)
 | 
|  |   1254 |     defer
 | 
|  |   1255 |     apply (rule rotate_conv_mod [symmetric])+
 | 
|  |   1256 |   apply simp
 | 
|  |   1257 |   done
 | 
|  |   1258 | 
 | 
|  |   1259 | lemmas rotate_eqs [standard] = 
 | 
|  |   1260 |   trans [OF rotate0 [THEN fun_cong] id_apply]
 | 
|  |   1261 |   rotate_rotate [symmetric] 
 | 
|  |   1262 |   rotate_id 
 | 
|  |   1263 |   rotate_conv_mod 
 | 
|  |   1264 |   rotate_eq_mod
 | 
|  |   1265 | 
 | 
|  |   1266 | 
 | 
| 24350 |   1267 | subsubsection "Rotation of list to right"
 | 
| 24333 |   1268 | 
 | 
|  |   1269 | lemma rotate1_rl': "rotater1 (l @ [a]) = a # l"
 | 
|  |   1270 |   unfolding rotater1_def by (cases l) auto
 | 
|  |   1271 | 
 | 
|  |   1272 | lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l"
 | 
|  |   1273 |   apply (unfold rotater1_def)
 | 
|  |   1274 |   apply (cases "l")
 | 
|  |   1275 |   apply (case_tac [2] "list")
 | 
|  |   1276 |   apply auto
 | 
|  |   1277 |   done
 | 
|  |   1278 | 
 | 
|  |   1279 | lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l"
 | 
|  |   1280 |   unfolding rotater1_def by (cases l) auto
 | 
|  |   1281 | 
 | 
|  |   1282 | lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)"
 | 
|  |   1283 |   apply (cases "xs")
 | 
|  |   1284 |   apply (simp add : rotater1_def)
 | 
|  |   1285 |   apply (simp add : rotate1_rl')
 | 
|  |   1286 |   done
 | 
|  |   1287 |   
 | 
|  |   1288 | lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)"
 | 
|  |   1289 |   unfolding rotater_def by (induct n) (auto intro: rotater1_rev')
 | 
|  |   1290 | 
 | 
|  |   1291 | lemmas rotater_rev = rotater_rev' [where xs = "rev ?ys", simplified]
 | 
|  |   1292 | 
 | 
|  |   1293 | lemma rotater_drop_take: 
 | 
|  |   1294 |   "rotater n xs = 
 | 
|  |   1295 |    drop (length xs - n mod length xs) xs @
 | 
|  |   1296 |    take (length xs - n mod length xs) xs"
 | 
|  |   1297 |   by (clarsimp simp add : rotater_rev rotate_drop_take rev_take rev_drop)
 | 
|  |   1298 | 
 | 
|  |   1299 | lemma rotater_Suc [simp] : 
 | 
|  |   1300 |   "rotater (Suc n) xs = rotater1 (rotater n xs)"
 | 
|  |   1301 |   unfolding rotater_def by auto
 | 
|  |   1302 | 
 | 
|  |   1303 | lemma rotate_inv_plus [rule_format] :
 | 
|  |   1304 |   "ALL k. k = m + n --> rotater k (rotate n xs) = rotater m xs & 
 | 
|  |   1305 |     rotate k (rotater n xs) = rotate m xs & 
 | 
|  |   1306 |     rotater n (rotate k xs) = rotate m xs & 
 | 
|  |   1307 |     rotate n (rotater k xs) = rotater m xs"
 | 
|  |   1308 |   unfolding rotater_def rotate_def
 | 
|  |   1309 |   by (induct n) (auto intro: funpow_swap1 [THEN trans])
 | 
|  |   1310 |   
 | 
|  |   1311 | lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus]
 | 
|  |   1312 | 
 | 
|  |   1313 | lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified]
 | 
|  |   1314 | 
 | 
|  |   1315 | lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1, standard]
 | 
|  |   1316 | lemmas rotate_rl [simp] =
 | 
|  |   1317 |   rotate_inv_eq [THEN conjunct2, THEN conjunct1, standard]
 | 
|  |   1318 | 
 | 
|  |   1319 | lemma rotate_gal: "(rotater n xs = ys) = (rotate n ys = xs)"
 | 
|  |   1320 |   by auto
 | 
|  |   1321 | 
 | 
|  |   1322 | lemma rotate_gal': "(ys = rotater n xs) = (xs = rotate n ys)"
 | 
|  |   1323 |   by auto
 | 
|  |   1324 | 
 | 
|  |   1325 | lemma length_rotater [simp]: 
 | 
|  |   1326 |   "length (rotater n xs) = length xs"
 | 
|  |   1327 |   by (simp add : rotater_rev)
 | 
|  |   1328 | 
 | 
|  |   1329 | lemmas rrs0 = rotate_eqs [THEN restrict_to_left, 
 | 
|  |   1330 |   simplified rotate_gal [symmetric] rotate_gal' [symmetric], standard]
 | 
|  |   1331 | lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]]
 | 
|  |   1332 | lemmas rotater_eqs = rrs1 [simplified length_rotater, standard]
 | 
|  |   1333 | lemmas rotater_0 = rotater_eqs (1)
 | 
|  |   1334 | lemmas rotater_add = rotater_eqs (2)
 | 
|  |   1335 | 
 | 
|  |   1336 | 
 | 
| 24350 |   1337 | subsubsection "map, app2, commuting with rotate(r)"
 | 
| 24333 |   1338 | 
 | 
|  |   1339 | lemma last_map: "xs ~= [] ==> last (map f xs) = f (last xs)"
 | 
|  |   1340 |   by (induct xs) auto
 | 
|  |   1341 | 
 | 
|  |   1342 | lemma butlast_map:
 | 
|  |   1343 |   "xs ~= [] ==> butlast (map f xs) = map f (butlast xs)"
 | 
|  |   1344 |   by (induct xs) auto
 | 
|  |   1345 | 
 | 
|  |   1346 | lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" 
 | 
|  |   1347 |   unfolding rotater1_def
 | 
|  |   1348 |   by (cases xs) (auto simp add: last_map butlast_map)
 | 
|  |   1349 | 
 | 
|  |   1350 | lemma rotater_map:
 | 
|  |   1351 |   "rotater n (map f xs) = map f (rotater n xs)" 
 | 
|  |   1352 |   unfolding rotater_def
 | 
|  |   1353 |   by (induct n) (auto simp add : rotater1_map)
 | 
|  |   1354 | 
 | 
|  |   1355 | lemma but_last_zip [rule_format] :
 | 
|  |   1356 |   "ALL ys. length xs = length ys --> xs ~= [] --> 
 | 
|  |   1357 |     last (zip xs ys) = (last xs, last ys) & 
 | 
|  |   1358 |     butlast (zip xs ys) = zip (butlast xs) (butlast ys)" 
 | 
|  |   1359 |   apply (induct "xs")
 | 
|  |   1360 |   apply auto
 | 
|  |   1361 |      apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+
 | 
|  |   1362 |   done
 | 
|  |   1363 | 
 | 
|  |   1364 | lemma but_last_app2 [rule_format] :
 | 
|  |   1365 |   "ALL ys. length xs = length ys --> xs ~= [] --> 
 | 
|  |   1366 |     last (app2 f xs ys) = f (last xs) (last ys) & 
 | 
|  |   1367 |     butlast (app2 f xs ys) = app2 f (butlast xs) (butlast ys)" 
 | 
|  |   1368 |   apply (induct "xs")
 | 
|  |   1369 |   apply auto
 | 
|  |   1370 |      apply (unfold app2_def)
 | 
|  |   1371 |      apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+
 | 
|  |   1372 |   done
 | 
|  |   1373 | 
 | 
|  |   1374 | lemma rotater1_zip:
 | 
|  |   1375 |   "length xs = length ys ==> 
 | 
|  |   1376 |     rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" 
 | 
|  |   1377 |   apply (unfold rotater1_def)
 | 
|  |   1378 |   apply (cases "xs")
 | 
|  |   1379 |    apply auto
 | 
|  |   1380 |    apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+
 | 
|  |   1381 |   done
 | 
|  |   1382 | 
 | 
|  |   1383 | lemma rotater1_app2:
 | 
|  |   1384 |   "length xs = length ys ==> 
 | 
|  |   1385 |     rotater1 (app2 f xs ys) = app2 f (rotater1 xs) (rotater1 ys)" 
 | 
|  |   1386 |   unfolding app2_def by (simp add: rotater1_map rotater1_zip)
 | 
|  |   1387 | 
 | 
|  |   1388 | lemmas lrth = 
 | 
|  |   1389 |   box_equals [OF asm_rl length_rotater [symmetric] 
 | 
|  |   1390 |                  length_rotater [symmetric], 
 | 
|  |   1391 |               THEN rotater1_app2]
 | 
|  |   1392 | 
 | 
|  |   1393 | lemma rotater_app2: 
 | 
|  |   1394 |   "length xs = length ys ==> 
 | 
|  |   1395 |     rotater n (app2 f xs ys) = app2 f (rotater n xs) (rotater n ys)" 
 | 
|  |   1396 |   by (induct n) (auto intro!: lrth)
 | 
|  |   1397 | 
 | 
|  |   1398 | lemma rotate1_app2:
 | 
|  |   1399 |   "length xs = length ys ==> 
 | 
|  |   1400 |     rotate1 (app2 f xs ys) = app2 f (rotate1 xs) (rotate1 ys)" 
 | 
|  |   1401 |   apply (unfold app2_def)
 | 
|  |   1402 |   apply (cases xs)
 | 
|  |   1403 |    apply (cases ys, auto simp add : rotate1_def)+
 | 
|  |   1404 |   done
 | 
|  |   1405 | 
 | 
|  |   1406 | lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] 
 | 
|  |   1407 |   length_rotate [symmetric], THEN rotate1_app2]
 | 
|  |   1408 | 
 | 
|  |   1409 | lemma rotate_app2: 
 | 
|  |   1410 |   "length xs = length ys ==> 
 | 
|  |   1411 |     rotate n (app2 f xs ys) = app2 f (rotate n xs) (rotate n ys)" 
 | 
|  |   1412 |   by (induct n) (auto intro!: lth)
 | 
|  |   1413 | 
 | 
|  |   1414 | 
 | 
|  |   1415 | -- "corresponding equalities for word rotation"
 | 
|  |   1416 | 
 | 
|  |   1417 | lemma to_bl_rotl: 
 | 
|  |   1418 |   "to_bl (word_rotl n w) = rotate n (to_bl w)"
 | 
|  |   1419 |   by (simp add: word_bl.Abs_inverse' word_rotl_def)
 | 
|  |   1420 | 
 | 
|  |   1421 | lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]]
 | 
|  |   1422 | 
 | 
|  |   1423 | lemmas word_rotl_eqs =
 | 
|  |   1424 |   blrs0 [simplified word_bl.Rep' word_bl.Rep_inject to_bl_rotl [symmetric]]
 | 
|  |   1425 | 
 | 
|  |   1426 | lemma to_bl_rotr: 
 | 
|  |   1427 |   "to_bl (word_rotr n w) = rotater n (to_bl w)"
 | 
|  |   1428 |   by (simp add: word_bl.Abs_inverse' word_rotr_def)
 | 
|  |   1429 | 
 | 
|  |   1430 | lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]]
 | 
|  |   1431 | 
 | 
|  |   1432 | lemmas word_rotr_eqs =
 | 
|  |   1433 |   brrs0 [simplified word_bl.Rep' word_bl.Rep_inject to_bl_rotr [symmetric]]
 | 
|  |   1434 | 
 | 
|  |   1435 | declare word_rotr_eqs (1) [simp]
 | 
|  |   1436 | declare word_rotl_eqs (1) [simp]
 | 
|  |   1437 | 
 | 
|  |   1438 | lemma
 | 
|  |   1439 |   word_rot_rl [simp]:
 | 
|  |   1440 |   "word_rotl k (word_rotr k v) = v" and
 | 
|  |   1441 |   word_rot_lr [simp]:
 | 
|  |   1442 |   "word_rotr k (word_rotl k v) = v"
 | 
|  |   1443 |   by (auto simp add: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric])
 | 
|  |   1444 | 
 | 
|  |   1445 | lemma
 | 
|  |   1446 |   word_rot_gal:
 | 
|  |   1447 |   "(word_rotr n v = w) = (word_rotl n w = v)" and
 | 
|  |   1448 |   word_rot_gal':
 | 
|  |   1449 |   "(w = word_rotr n v) = (v = word_rotl n w)"
 | 
|  |   1450 |   by (auto simp: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric] 
 | 
|  |   1451 |            dest: sym)
 | 
|  |   1452 | 
 | 
|  |   1453 | lemma word_rotr_rev:
 | 
|  |   1454 |   "word_rotr n w = word_reverse (word_rotl n (word_reverse w))"
 | 
|  |   1455 |   by (simp add: word_bl.Rep_inject [symmetric] to_bl_word_rev
 | 
|  |   1456 |                 to_bl_rotr to_bl_rotl rotater_rev)
 | 
|  |   1457 |   
 | 
|  |   1458 | lemma word_roti_0 [simp]: "word_roti 0 w = w"
 | 
|  |   1459 |   by (unfold word_rot_defs) auto
 | 
|  |   1460 | 
 | 
|  |   1461 | lemmas abl_cong = arg_cong [where f = "of_bl"]
 | 
|  |   1462 | 
 | 
|  |   1463 | lemma word_roti_add: 
 | 
|  |   1464 |   "word_roti (m + n) w = word_roti m (word_roti n w)"
 | 
|  |   1465 | proof -
 | 
|  |   1466 |   have rotater_eq_lem: 
 | 
|  |   1467 |     "\<And>m n xs. m = n ==> rotater m xs = rotater n xs"
 | 
|  |   1468 |     by auto
 | 
|  |   1469 | 
 | 
|  |   1470 |   have rotate_eq_lem: 
 | 
|  |   1471 |     "\<And>m n xs. m = n ==> rotate m xs = rotate n xs"
 | 
|  |   1472 |     by auto
 | 
|  |   1473 | 
 | 
|  |   1474 |   note rpts [symmetric, standard] = 
 | 
|  |   1475 |     rotate_inv_plus [THEN conjunct1]
 | 
|  |   1476 |     rotate_inv_plus [THEN conjunct2, THEN conjunct1]
 | 
|  |   1477 |     rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct1]
 | 
|  |   1478 |     rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct2]
 | 
|  |   1479 | 
 | 
|  |   1480 |   note rrp = trans [symmetric, OF rotate_rotate rotate_eq_lem]
 | 
|  |   1481 |   note rrrp = trans [symmetric, OF rotater_add [symmetric] rotater_eq_lem]
 | 
|  |   1482 | 
 | 
|  |   1483 |   show ?thesis
 | 
|  |   1484 |   apply (unfold word_rot_defs)
 | 
|  |   1485 |   apply (simp only: split: split_if)
 | 
|  |   1486 |   apply (safe intro!: abl_cong)
 | 
|  |   1487 |   apply (simp_all only: to_bl_rotl [THEN word_bl.Rep_inverse'] 
 | 
|  |   1488 |                     to_bl_rotl
 | 
|  |   1489 |                     to_bl_rotr [THEN word_bl.Rep_inverse']
 | 
|  |   1490 |                     to_bl_rotr)
 | 
|  |   1491 |   apply (rule rrp rrrp rpts,
 | 
|  |   1492 |          simp add: nat_add_distrib [symmetric] 
 | 
|  |   1493 |                    nat_diff_distrib [symmetric])+
 | 
|  |   1494 |   done
 | 
|  |   1495 | qed
 | 
|  |   1496 |     
 | 
|  |   1497 | lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w"
 | 
|  |   1498 |   apply (unfold word_rot_defs)
 | 
|  |   1499 |   apply (cut_tac y="size w" in gt_or_eq_0)
 | 
|  |   1500 |   apply (erule disjE)
 | 
|  |   1501 |    apply simp_all
 | 
|  |   1502 |   apply (safe intro!: abl_cong)
 | 
|  |   1503 |    apply (rule rotater_eqs)
 | 
|  |   1504 |    apply (simp add: word_size nat_mod_distrib)
 | 
|  |   1505 |   apply (simp add: rotater_add [symmetric] rotate_gal [symmetric])
 | 
|  |   1506 |   apply (rule rotater_eqs)
 | 
|  |   1507 |   apply (simp add: word_size nat_mod_distrib)
 | 
|  |   1508 |   apply (rule int_eq_0_conv [THEN iffD1])
 | 
|  |   1509 |   apply (simp only: zmod_int zadd_int [symmetric])
 | 
|  |   1510 |   apply (simp add: rdmods)
 | 
|  |   1511 |   done
 | 
|  |   1512 | 
 | 
|  |   1513 | lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size]
 | 
|  |   1514 | 
 | 
|  |   1515 | 
 | 
| 24350 |   1516 | subsubsection "Word rotation commutes with bit-wise operations"
 | 
| 24333 |   1517 | 
 | 
|  |   1518 | (* using locale to not pollute lemma namespace *)
 | 
|  |   1519 | locale word_rotate 
 | 
|  |   1520 | 
 | 
|  |   1521 | context word_rotate
 | 
|  |   1522 | begin
 | 
|  |   1523 | 
 | 
|  |   1524 | lemmas word_rot_defs' = to_bl_rotl to_bl_rotr
 | 
|  |   1525 | 
 | 
|  |   1526 | lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor
 | 
|  |   1527 | 
 | 
|  |   1528 | lemmas lbl_lbl = trans [OF word_bl.Rep' word_bl.Rep' [symmetric]]
 | 
|  |   1529 | 
 | 
|  |   1530 | lemmas ths_app2 [OF lbl_lbl] = rotate_app2 rotater_app2
 | 
|  |   1531 | 
 | 
|  |   1532 | lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map
 | 
|  |   1533 | 
 | 
|  |   1534 | lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_app2 ths_map
 | 
|  |   1535 | 
 | 
|  |   1536 | lemma word_rot_logs:
 | 
|  |   1537 |   "word_rotl n (NOT v) = NOT word_rotl n v"
 | 
|  |   1538 |   "word_rotr n (NOT v) = NOT word_rotr n v"
 | 
|  |   1539 |   "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y"
 | 
|  |   1540 |   "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y"
 | 
|  |   1541 |   "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y"
 | 
|  |   1542 |   "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y"
 | 
|  |   1543 |   "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y"
 | 
|  |   1544 |   "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y"  
 | 
|  |   1545 |   by (rule word_bl.Rep_eqD,
 | 
|  |   1546 |       rule word_rot_defs' [THEN trans],
 | 
|  |   1547 |       simp only: blwl_syms [symmetric],
 | 
|  |   1548 |       rule th1s [THEN trans], 
 | 
|  |   1549 |       rule refl)+
 | 
|  |   1550 | end
 | 
|  |   1551 | 
 | 
|  |   1552 | lemmas word_rot_logs = word_rotate.word_rot_logs
 | 
|  |   1553 | 
 | 
|  |   1554 | lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take,
 | 
|  |   1555 |   simplified word_bl.Rep', standard]
 | 
|  |   1556 | 
 | 
|  |   1557 | lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take,
 | 
|  |   1558 |   simplified word_bl.Rep', standard]
 | 
|  |   1559 | 
 | 
|  |   1560 | lemma bl_word_roti_dt': 
 | 
| 24465 |   1561 |   "n = nat ((- i) mod int (size (w :: 'a :: len word))) ==> 
 | 
| 24333 |   1562 |     to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)"
 | 
|  |   1563 |   apply (unfold word_roti_def)
 | 
|  |   1564 |   apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size)
 | 
|  |   1565 |   apply safe
 | 
|  |   1566 |    apply (simp add: zmod_zminus1_eq_if)
 | 
|  |   1567 |    apply safe
 | 
|  |   1568 |     apply (simp add: nat_mult_distrib)
 | 
|  |   1569 |    apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj 
 | 
|  |   1570 |                                       [THEN conjunct2, THEN order_less_imp_le]]
 | 
|  |   1571 |                     nat_mod_distrib)
 | 
|  |   1572 |   apply (simp add: nat_mod_distrib)
 | 
|  |   1573 |   done
 | 
|  |   1574 | 
 | 
|  |   1575 | lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size]
 | 
|  |   1576 | 
 | 
|  |   1577 | lemmas word_rotl_dt = bl_word_rotl_dt 
 | 
|  |   1578 |   [THEN word_bl.Rep_inverse' [symmetric], standard] 
 | 
|  |   1579 | lemmas word_rotr_dt = bl_word_rotr_dt 
 | 
|  |   1580 |   [THEN word_bl.Rep_inverse' [symmetric], standard]
 | 
|  |   1581 | lemmas word_roti_dt = bl_word_roti_dt 
 | 
|  |   1582 |   [THEN word_bl.Rep_inverse' [symmetric], standard]
 | 
|  |   1583 | 
 | 
|  |   1584 | lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 & word_rotl i 0 = 0"
 | 
|  |   1585 |   by (simp add : word_rotr_dt word_rotl_dt to_bl_0 replicate_add [symmetric])
 | 
|  |   1586 | 
 | 
|  |   1587 | lemma word_roti_0' [simp] : "word_roti n 0 = 0"
 | 
|  |   1588 |   unfolding word_roti_def by auto
 | 
|  |   1589 | 
 | 
|  |   1590 | lemmas word_rotr_dt_no_bin' [simp] = 
 | 
|  |   1591 |   word_rotr_dt [where w="number_of ?w", unfolded to_bl_no_bin]
 | 
|  |   1592 | 
 | 
|  |   1593 | lemmas word_rotl_dt_no_bin' [simp] = 
 | 
|  |   1594 |   word_rotl_dt [where w="number_of ?w", unfolded to_bl_no_bin]
 | 
|  |   1595 | 
 | 
|  |   1596 | declare word_roti_def [simp]
 | 
|  |   1597 | 
 | 
|  |   1598 | end
 | 
|  |   1599 | 
 |