src/HOL/Word/Reversed_Bit_Lists.thy
changeset 72515 c7038c397ae3
parent 72514 d8661799afb2
child 72516 17dc99589a91
child 72528 c435a4750636
equal deleted inserted replaced
72514:d8661799afb2 72515:c7038c397ae3
     1 (*  Title:      HOL/Word/Reversed_Bit_Lists.thy
       
     2     Author:     Jeremy Dawson, NICTA
       
     3 *)
       
     4 
       
     5 section \<open>Bit values as reversed lists of bools\<close>
       
     6 
       
     7 theory Reversed_Bit_Lists
       
     8   imports Word Misc_Typedef
       
     9 begin
       
    10 
       
    11 context comm_semiring_1
       
    12 begin
       
    13 
       
    14 lemma horner_sum_append:
       
    15   \<open>horner_sum f a (xs @ ys) = horner_sum f a xs + a ^ length xs * horner_sum f a ys\<close>
       
    16   using sum.atLeastLessThan_shift_bounds [of _ 0 \<open>length xs\<close> \<open>length ys\<close>]
       
    17     atLeastLessThan_add_Un [of 0 \<open>length xs\<close> \<open>length ys\<close>]
       
    18   by (simp add: horner_sum_eq_sum sum_distrib_left sum.union_disjoint ac_simps nth_append power_add)
       
    19 
       
    20 end
       
    21 
       
    22 lemma horner_sum_of_bool_2_concat:
       
    23   \<open>horner_sum of_bool 2 (concat (map (\<lambda>x. map (bit x) [0..<LENGTH('a)]) ws)) = horner_sum uint (2 ^ LENGTH('a)) ws\<close>
       
    24   for ws :: \<open>'a::len word list\<close>
       
    25 proof (induction ws)
       
    26   case Nil
       
    27   then show ?case
       
    28     by simp
       
    29 next
       
    30   case (Cons w ws)
       
    31   moreover have \<open>horner_sum of_bool 2 (map (bit w) [0..<LENGTH('a)]) = uint w\<close>
       
    32   proof transfer
       
    33     fix k :: int
       
    34     have \<open>map (\<lambda>n. n < LENGTH('a) \<and> bit k n) [0..<LENGTH('a)]
       
    35       = map (bit k) [0..<LENGTH('a)]\<close>
       
    36       by simp
       
    37     then show \<open>horner_sum of_bool 2 (map (\<lambda>n. n < LENGTH('a) \<and> bit k n) [0..<LENGTH('a)])
       
    38       = take_bit LENGTH('a) k\<close>
       
    39       by (simp only: horner_sum_bit_eq_take_bit)
       
    40   qed
       
    41   ultimately show ?case
       
    42     by (simp add: horner_sum_append)
       
    43 qed
       
    44 
       
    45 
       
    46 subsection \<open>Implicit augmentation of list prefixes\<close>
       
    47 
       
    48 primrec takefill :: "'a \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
       
    49 where
       
    50     Z: "takefill fill 0 xs = []"
       
    51   | Suc: "takefill fill (Suc n) xs =
       
    52       (case xs of
       
    53         [] \<Rightarrow> fill # takefill fill n xs
       
    54       | y # ys \<Rightarrow> y # takefill fill n ys)"
       
    55 
       
    56 lemma nth_takefill: "m < n \<Longrightarrow> takefill fill n l ! m = (if m < length l then l ! m else fill)"
       
    57   apply (induct n arbitrary: m l)
       
    58    apply clarsimp
       
    59   apply clarsimp
       
    60   apply (case_tac m)
       
    61    apply (simp split: list.split)
       
    62   apply (simp split: list.split)
       
    63   done
       
    64 
       
    65 lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill"
       
    66   by (induct n arbitrary: l) (auto split: list.split)
       
    67 
       
    68 lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill"
       
    69   by (simp add: takefill_alt replicate_add [symmetric])
       
    70 
       
    71 lemma takefill_le': "n = m + k \<Longrightarrow> takefill x m (takefill x n l) = takefill x m l"
       
    72   by (induct m arbitrary: l n) (auto split: list.split)
       
    73 
       
    74 lemma length_takefill [simp]: "length (takefill fill n l) = n"
       
    75   by (simp add: takefill_alt)
       
    76 
       
    77 lemma take_takefill': "n = k + m \<Longrightarrow> take k (takefill fill n w) = takefill fill k w"
       
    78   by (induct k arbitrary: w n) (auto split: list.split)
       
    79 
       
    80 lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)"
       
    81   by (induct k arbitrary: w) (auto split: list.split)
       
    82 
       
    83 lemma takefill_le [simp]: "m \<le> n \<Longrightarrow> takefill x m (takefill x n l) = takefill x m l"
       
    84   by (auto simp: le_iff_add takefill_le')
       
    85 
       
    86 lemma take_takefill [simp]: "m \<le> n \<Longrightarrow> take m (takefill fill n w) = takefill fill m w"
       
    87   by (auto simp: le_iff_add take_takefill')
       
    88 
       
    89 lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)"
       
    90   by (induct xs) auto
       
    91 
       
    92 lemma takefill_same': "l = length xs \<Longrightarrow> takefill fill l xs = xs"
       
    93   by (induct xs arbitrary: l) auto
       
    94 
       
    95 lemmas takefill_same [simp] = takefill_same' [OF refl]
       
    96 
       
    97 lemma tf_rev:
       
    98   "n + k = m + length bl \<Longrightarrow> takefill x m (rev (takefill y n bl)) =
       
    99     rev (takefill y m (rev (takefill x k (rev bl))))"
       
   100   apply (rule nth_equalityI)
       
   101    apply (auto simp add: nth_takefill rev_nth)
       
   102   apply (rule_tac f = "\<lambda>n. bl ! n" in arg_cong)
       
   103   apply arith
       
   104   done
       
   105 
       
   106 lemma takefill_minus: "0 < n \<Longrightarrow> takefill fill (Suc (n - 1)) w = takefill fill n w"
       
   107   by auto
       
   108 
       
   109 lemmas takefill_Suc_cases =
       
   110   list.cases [THEN takefill.Suc [THEN trans]]
       
   111 
       
   112 lemmas takefill_Suc_Nil = takefill_Suc_cases (1)
       
   113 lemmas takefill_Suc_Cons = takefill_Suc_cases (2)
       
   114 
       
   115 lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2]
       
   116   takefill_minus [symmetric, THEN trans]]
       
   117 
       
   118 lemma takefill_numeral_Nil [simp]:
       
   119   "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []"
       
   120   by (simp add: numeral_eq_Suc)
       
   121 
       
   122 lemma takefill_numeral_Cons [simp]:
       
   123   "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs"
       
   124   by (simp add: numeral_eq_Suc)
       
   125 
       
   126 
       
   127 subsection \<open>Range projection\<close>
       
   128 
       
   129 definition bl_of_nth :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> 'a list"
       
   130   where "bl_of_nth n f = map f (rev [0..<n])"
       
   131 
       
   132 lemma bl_of_nth_simps [simp, code]:
       
   133   "bl_of_nth 0 f = []"
       
   134   "bl_of_nth (Suc n) f = f n # bl_of_nth n f"
       
   135   by (simp_all add: bl_of_nth_def)
       
   136 
       
   137 lemma length_bl_of_nth [simp]: "length (bl_of_nth n f) = n"
       
   138   by (simp add: bl_of_nth_def)
       
   139 
       
   140 lemma nth_bl_of_nth [simp]: "m < n \<Longrightarrow> rev (bl_of_nth n f) ! m = f m"
       
   141   by (simp add: bl_of_nth_def rev_map)
       
   142 
       
   143 lemma bl_of_nth_inj: "(\<And>k. k < n \<Longrightarrow> f k = g k) \<Longrightarrow> bl_of_nth n f = bl_of_nth n g"
       
   144   by (simp add: bl_of_nth_def)
       
   145 
       
   146 lemma bl_of_nth_nth_le: "n \<le> length xs \<Longrightarrow> bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs"
       
   147   apply (induct n arbitrary: xs)
       
   148    apply clarsimp
       
   149   apply clarsimp
       
   150   apply (rule trans [OF _ hd_Cons_tl])
       
   151    apply (frule Suc_le_lessD)
       
   152    apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric])
       
   153    apply (subst hd_drop_conv_nth)
       
   154     apply force
       
   155    apply simp_all
       
   156   apply (rule_tac f = "\<lambda>n. drop n xs" in arg_cong)
       
   157   apply simp
       
   158   done
       
   159 
       
   160 lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs"
       
   161   by (simp add: bl_of_nth_nth_le)
       
   162 
       
   163 
       
   164 subsection \<open>More\<close>
       
   165 
       
   166 definition rotater1 :: "'a list \<Rightarrow> 'a list"
       
   167   where "rotater1 ys =
       
   168     (case ys of [] \<Rightarrow> [] | x # xs \<Rightarrow> last ys # butlast ys)"
       
   169 
       
   170 definition rotater :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
       
   171   where "rotater n = rotater1 ^^ n"
       
   172 
       
   173 lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified]
       
   174 
       
   175 lemma rotate1_rl': "rotater1 (l @ [a]) = a # l"
       
   176   by (cases l) (auto simp: rotater1_def)
       
   177 
       
   178 lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l"
       
   179   apply (unfold rotater1_def)
       
   180   apply (cases "l")
       
   181    apply (case_tac [2] "list")
       
   182     apply auto
       
   183   done
       
   184 
       
   185 lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l"
       
   186   by (cases l) (auto simp: rotater1_def)
       
   187 
       
   188 lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)"
       
   189   by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl')
       
   190 
       
   191 lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)"
       
   192   by (induct n) (auto simp: rotater_def intro: rotater1_rev')
       
   193 
       
   194 lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))"
       
   195   using rotater_rev' [where xs = "rev ys"] by simp
       
   196 
       
   197 lemma rotater_drop_take:
       
   198   "rotater n xs =
       
   199     drop (length xs - n mod length xs) xs @
       
   200     take (length xs - n mod length xs) xs"
       
   201   by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop)
       
   202 
       
   203 lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)"
       
   204   unfolding rotater_def by auto
       
   205 
       
   206 lemma nth_rotater:
       
   207   \<open>rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\<close> if \<open>n < length xs\<close>
       
   208   using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq)
       
   209 
       
   210 lemma nth_rotater1:
       
   211   \<open>rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\<close> if \<open>n < length xs\<close>
       
   212   using that nth_rotater [of n xs 1] by simp
       
   213 
       
   214 lemma rotate_inv_plus [rule_format]:
       
   215   "\<forall>k. k = m + n \<longrightarrow> rotater k (rotate n xs) = rotater m xs \<and>
       
   216     rotate k (rotater n xs) = rotate m xs \<and>
       
   217     rotater n (rotate k xs) = rotate m xs \<and>
       
   218     rotate n (rotater k xs) = rotater m xs"
       
   219   by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans])
       
   220 
       
   221 lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus]
       
   222 
       
   223 lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified]
       
   224 
       
   225 lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1]
       
   226 lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1]
       
   227 
       
   228 lemma rotate_gal: "rotater n xs = ys \<longleftrightarrow> rotate n ys = xs"
       
   229   by auto
       
   230 
       
   231 lemma rotate_gal': "ys = rotater n xs \<longleftrightarrow> xs = rotate n ys"
       
   232   by auto
       
   233 
       
   234 lemma length_rotater [simp]: "length (rotater n xs) = length xs"
       
   235   by (simp add : rotater_rev)
       
   236 
       
   237 lemma rotate_eq_mod: "m mod length xs = n mod length xs \<Longrightarrow> rotate m xs = rotate n xs"
       
   238   apply (rule box_equals)
       
   239     defer
       
   240     apply (rule rotate_conv_mod [symmetric])+
       
   241   apply simp
       
   242   done
       
   243 
       
   244 lemma restrict_to_left: "x = y \<Longrightarrow> x = z \<longleftrightarrow> y = z"
       
   245   by simp
       
   246 
       
   247 lemmas rotate_eqs =
       
   248   trans [OF rotate0 [THEN fun_cong] id_apply]
       
   249   rotate_rotate [symmetric]
       
   250   rotate_id
       
   251   rotate_conv_mod
       
   252   rotate_eq_mod
       
   253 
       
   254 lemmas rrs0 = rotate_eqs [THEN restrict_to_left,
       
   255   simplified rotate_gal [symmetric] rotate_gal' [symmetric]]
       
   256 lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]]
       
   257 lemmas rotater_eqs = rrs1 [simplified length_rotater]
       
   258 lemmas rotater_0 = rotater_eqs (1)
       
   259 lemmas rotater_add = rotater_eqs (2)
       
   260 
       
   261 lemma butlast_map: "xs \<noteq> [] \<Longrightarrow> butlast (map f xs) = map f (butlast xs)"
       
   262   by (induct xs) auto
       
   263 
       
   264 lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)"
       
   265   by (cases xs) (auto simp: rotater1_def last_map butlast_map)
       
   266 
       
   267 lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)"
       
   268   by (induct n) (auto simp: rotater_def rotater1_map)
       
   269 
       
   270 lemma but_last_zip [rule_format] :
       
   271   "\<forall>ys. length xs = length ys \<longrightarrow> xs \<noteq> [] \<longrightarrow>
       
   272     last (zip xs ys) = (last xs, last ys) \<and>
       
   273     butlast (zip xs ys) = zip (butlast xs) (butlast ys)"
       
   274   apply (induct xs)
       
   275    apply auto
       
   276      apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+
       
   277   done
       
   278 
       
   279 lemma but_last_map2 [rule_format] :
       
   280   "\<forall>ys. length xs = length ys \<longrightarrow> xs \<noteq> [] \<longrightarrow>
       
   281     last (map2 f xs ys) = f (last xs) (last ys) \<and>
       
   282     butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)"
       
   283   apply (induct xs)
       
   284    apply auto
       
   285      apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+
       
   286   done
       
   287 
       
   288 lemma rotater1_zip:
       
   289   "length xs = length ys \<Longrightarrow>
       
   290     rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)"
       
   291   apply (unfold rotater1_def)
       
   292   apply (cases xs)
       
   293    apply auto
       
   294    apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+
       
   295   done
       
   296 
       
   297 lemma rotater1_map2:
       
   298   "length xs = length ys \<Longrightarrow>
       
   299     rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)"
       
   300   by (simp add: rotater1_map rotater1_zip)
       
   301 
       
   302 lemmas lrth =
       
   303   box_equals [OF asm_rl length_rotater [symmetric]
       
   304                  length_rotater [symmetric],
       
   305               THEN rotater1_map2]
       
   306 
       
   307 lemma rotater_map2:
       
   308   "length xs = length ys \<Longrightarrow>
       
   309     rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)"
       
   310   by (induct n) (auto intro!: lrth)
       
   311 
       
   312 lemma rotate1_map2:
       
   313   "length xs = length ys \<Longrightarrow>
       
   314     rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)"
       
   315   by (cases xs; cases ys) auto
       
   316 
       
   317 lemmas lth = box_equals [OF asm_rl length_rotate [symmetric]
       
   318   length_rotate [symmetric], THEN rotate1_map2]
       
   319 
       
   320 lemma rotate_map2:
       
   321   "length xs = length ys \<Longrightarrow>
       
   322     rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)"
       
   323   by (induct n) (auto intro!: lth)
       
   324 
       
   325 
       
   326 subsection \<open>Explicit bit representation of \<^typ>\<open>int\<close>\<close>
       
   327 
       
   328 primrec bl_to_bin_aux :: "bool list \<Rightarrow> int \<Rightarrow> int"
       
   329   where
       
   330     Nil: "bl_to_bin_aux [] w = w"
       
   331   | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)"
       
   332 
       
   333 definition bl_to_bin :: "bool list \<Rightarrow> int"
       
   334   where "bl_to_bin bs = bl_to_bin_aux bs 0"
       
   335 
       
   336 primrec bin_to_bl_aux :: "nat \<Rightarrow> int \<Rightarrow> bool list \<Rightarrow> bool list"
       
   337   where
       
   338     Z: "bin_to_bl_aux 0 w bl = bl"
       
   339   | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)"
       
   340 
       
   341 definition bin_to_bl :: "nat \<Rightarrow> int \<Rightarrow> bool list"
       
   342   where "bin_to_bl n w = bin_to_bl_aux n w []"
       
   343 
       
   344 lemma bin_to_bl_aux_zero_minus_simp [simp]:
       
   345   "0 < n \<Longrightarrow> bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)"
       
   346   by (cases n) auto
       
   347 
       
   348 lemma bin_to_bl_aux_minus1_minus_simp [simp]:
       
   349   "0 < n \<Longrightarrow> bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)"
       
   350   by (cases n) auto
       
   351 
       
   352 lemma bin_to_bl_aux_one_minus_simp [simp]:
       
   353   "0 < n \<Longrightarrow> bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)"
       
   354   by (cases n) auto
       
   355 
       
   356 lemma bin_to_bl_aux_Bit0_minus_simp [simp]:
       
   357   "0 < n \<Longrightarrow>
       
   358     bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)"
       
   359   by (cases n) simp_all
       
   360 
       
   361 lemma bin_to_bl_aux_Bit1_minus_simp [simp]:
       
   362   "0 < n \<Longrightarrow>
       
   363     bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)"
       
   364   by (cases n) simp_all
       
   365 
       
   366 lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)"
       
   367   by (induct bs arbitrary: w) auto
       
   368 
       
   369 lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)"
       
   370   by (induct n arbitrary: w bs) auto
       
   371 
       
   372 lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)"
       
   373   unfolding bl_to_bin_def by (rule bl_to_bin_aux_append)
       
   374 
       
   375 lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs"
       
   376   by (simp add: bin_to_bl_def bin_to_bl_aux_append)
       
   377 
       
   378 lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []"
       
   379   by (auto simp: bin_to_bl_def)
       
   380 
       
   381 lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs"
       
   382   by (induct n arbitrary: w bs) auto
       
   383 
       
   384 lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n"
       
   385   by (simp add: bin_to_bl_def size_bin_to_bl_aux)
       
   386 
       
   387 lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs"
       
   388   apply (induct bs arbitrary: w n)
       
   389    apply auto
       
   390     apply (simp_all only: add_Suc [symmetric])
       
   391     apply (auto simp add: bin_to_bl_def)
       
   392   done
       
   393 
       
   394 lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs"
       
   395   unfolding bl_to_bin_def
       
   396   apply (rule box_equals)
       
   397     apply (rule bl_bin_bl')
       
   398    prefer 2
       
   399    apply (rule bin_to_bl_aux.Z)
       
   400   apply simp
       
   401   done
       
   402 
       
   403 lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \<Longrightarrow> length bs = length cs \<Longrightarrow> bs = cs"
       
   404   apply (rule_tac box_equals)
       
   405     defer
       
   406     apply (rule bl_bin_bl)
       
   407    apply (rule bl_bin_bl)
       
   408   apply simp
       
   409   done
       
   410 
       
   411 lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl"
       
   412   by (auto simp: bl_to_bin_def)
       
   413 
       
   414 lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0"
       
   415   by (auto simp: bl_to_bin_def)
       
   416 
       
   417 lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl"
       
   418   by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same)
       
   419 
       
   420 lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False"
       
   421   by (simp add: bin_to_bl_def bin_to_bl_zero_aux)
       
   422 
       
   423 lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl"
       
   424   by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same)
       
   425 
       
   426 lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True"
       
   427   by (simp add: bin_to_bl_def bin_to_bl_minus1_aux)
       
   428 
       
   429 
       
   430 subsection \<open>Semantic interpretation of \<^typ>\<open>bool list\<close> as \<^typ>\<open>int\<close>\<close>
       
   431 
       
   432 lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)"
       
   433   by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd)
       
   434 
       
   435 lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w"
       
   436   by (auto simp: bin_to_bl_def bin_bl_bin')
       
   437 
       
   438 lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl"
       
   439   by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def)
       
   440 
       
   441 lemma bin_to_bl_trunc [simp]: "n \<le> m \<Longrightarrow> bin_to_bl n (bintrunc m w) = bin_to_bl n w"
       
   442   by (auto intro: bl_to_bin_inj)
       
   443 
       
   444 lemma bin_to_bl_aux_bintr:
       
   445   "bin_to_bl_aux n (bintrunc m bin) bl =
       
   446     replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl"
       
   447   apply (induct n arbitrary: m bin bl)
       
   448    apply clarsimp
       
   449   apply clarsimp
       
   450   apply (case_tac "m")
       
   451    apply (clarsimp simp: bin_to_bl_zero_aux)
       
   452    apply (erule thin_rl)
       
   453    apply (induct_tac n)
       
   454     apply (auto simp add: take_bit_Suc)
       
   455   done
       
   456 
       
   457 lemma bin_to_bl_bintr:
       
   458   "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin"
       
   459   unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr)
       
   460 
       
   461 lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0"
       
   462   by (induct n) auto
       
   463 
       
   464 lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs"
       
   465   by (fact size_bin_to_bl_aux)
       
   466 
       
   467 lemma len_bin_to_bl: "length (bin_to_bl n w) = n"
       
   468   by (fact size_bin_to_bl) (* FIXME: duplicate *)
       
   469 
       
   470 lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w"
       
   471   by (induction bs arbitrary: w) (simp_all add: bin_sign_def)
       
   472 
       
   473 lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0"
       
   474   by (simp add: bl_to_bin_def sign_bl_bin')
       
   475 
       
   476 lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)"
       
   477   by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc)
       
   478 
       
   479 lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)"
       
   480   unfolding bin_to_bl_def by (rule bl_sbin_sign_aux)
       
   481 
       
   482 lemma bin_nth_of_bl_aux:
       
   483   "bin_nth (bl_to_bin_aux bl w) n =
       
   484     (n < size bl \<and> rev bl ! n \<or> n \<ge> length bl \<and> bin_nth w (n - size bl))"
       
   485   apply (induction bl arbitrary: w)
       
   486    apply simp_all
       
   487   apply safe
       
   488                       apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits)
       
   489   done
       
   490 
       
   491 lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \<and> rev bl ! n)"
       
   492   by (simp add: bl_to_bin_def bin_nth_of_bl_aux)
       
   493 
       
   494 lemma bin_nth_bl: "n < m \<Longrightarrow> bin_nth w n = nth (rev (bin_to_bl m w)) n"
       
   495   apply (induct n arbitrary: m w)
       
   496    apply clarsimp
       
   497    apply (case_tac m, clarsimp)
       
   498    apply (clarsimp simp: bin_to_bl_def)
       
   499    apply (simp add: bin_to_bl_aux_alt)
       
   500   apply (case_tac m, clarsimp)
       
   501   apply (clarsimp simp: bin_to_bl_def)
       
   502   apply (simp add: bin_to_bl_aux_alt bit_Suc)
       
   503   done
       
   504 
       
   505 lemma nth_bin_to_bl_aux:
       
   506   "n < m + length bl \<Longrightarrow> (bin_to_bl_aux m w bl) ! n =
       
   507     (if n < m then bit w (m - 1 - n) else bl ! (n - m))"
       
   508   apply (induction bl arbitrary: w)
       
   509    apply simp_all
       
   510    apply (simp add: bin_nth_bl [of \<open>m - Suc n\<close> m] rev_nth flip: bin_to_bl_def)
       
   511    apply (metis One_nat_def Suc_pred add_diff_cancel_left'
       
   512      add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def
       
   513      diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj
       
   514      less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux)
       
   515   done
       
   516 
       
   517 lemma nth_bin_to_bl: "n < m \<Longrightarrow> (bin_to_bl m w) ! n = bin_nth w (m - Suc n)"
       
   518   by (simp add: bin_to_bl_def nth_bin_to_bl_aux)
       
   519 
       
   520 lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))"
       
   521   apply (rule nth_equalityI)
       
   522    apply simp
       
   523   apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl)
       
   524   done
       
   525 
       
   526 lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))"
       
   527   by (simp add: takefill_bintrunc)
       
   528 
       
   529 lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)"
       
   530 proof (induction bs arbitrary: w)
       
   531   case Nil
       
   532   then show ?case
       
   533     by simp
       
   534 next
       
   535   case (Cons b bs)
       
   536   from Cons.IH [of \<open>1 + 2 * w\<close>] Cons.IH [of \<open>2 * w\<close>]
       
   537   show ?case
       
   538     apply (auto simp add: algebra_simps)
       
   539     apply (subst mult_2 [of \<open>2 ^ length bs\<close>])
       
   540     apply (simp only: add.assoc)
       
   541     apply (rule pos_add_strict)
       
   542      apply simp_all
       
   543     done
       
   544 qed
       
   545 
       
   546 lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)"
       
   547 proof (induct bs)
       
   548   case Nil
       
   549   then show ?case by simp
       
   550 next
       
   551   case (Cons b bs)
       
   552   with bl_to_bin_lt2p_aux[where w=1] show ?case
       
   553     by (simp add: bl_to_bin_def)
       
   554 qed
       
   555 
       
   556 lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs"
       
   557   by (metis bin_bl_bin bintr_lt2p bl_bin_bl)
       
   558 
       
   559 lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \<ge> w * (2 ^ length bs)"
       
   560 proof (induction bs arbitrary: w)
       
   561   case Nil
       
   562   then show ?case
       
   563     by simp
       
   564 next
       
   565   case (Cons b bs)
       
   566   from Cons.IH [of \<open>1 + 2 * w\<close>] Cons.IH [of \<open>2 * w\<close>]
       
   567   show ?case
       
   568     apply (auto simp add: algebra_simps)
       
   569     apply (rule add_le_imp_le_left [of \<open>2 ^ length bs\<close>])
       
   570     apply (rule add_increasing)
       
   571     apply simp_all
       
   572     done
       
   573 qed
       
   574 
       
   575 lemma bl_to_bin_ge0: "bl_to_bin bs \<ge> 0"
       
   576   apply (unfold bl_to_bin_def)
       
   577   apply (rule xtrans(4))
       
   578    apply (rule bl_to_bin_ge2p_aux)
       
   579   apply simp
       
   580   done
       
   581 
       
   582 lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)"
       
   583   apply (unfold bin_to_bl_def)
       
   584   apply (cases n, clarsimp)
       
   585   apply clarsimp
       
   586   apply (auto simp add: bin_to_bl_aux_alt)
       
   587   done
       
   588 
       
   589 lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))"
       
   590   using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp
       
   591 
       
   592 lemma butlast_rest_bl2bin_aux:
       
   593   "bl \<noteq> [] \<Longrightarrow> bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)"
       
   594   by (induct bl arbitrary: w) auto
       
   595 
       
   596 lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)"
       
   597   by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux)
       
   598 
       
   599 lemma trunc_bl2bin_aux:
       
   600   "bintrunc m (bl_to_bin_aux bl w) =
       
   601     bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)"
       
   602 proof (induct bl arbitrary: w)
       
   603   case Nil
       
   604   show ?case by simp
       
   605 next
       
   606   case (Cons b bl)
       
   607   show ?case
       
   608   proof (cases "m - length bl")
       
   609     case 0
       
   610     then have "Suc (length bl) - m = Suc (length bl - m)" by simp
       
   611     with Cons show ?thesis by simp
       
   612   next
       
   613     case (Suc n)
       
   614     then have "m - Suc (length bl) = n" by simp
       
   615     with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps)
       
   616   qed
       
   617 qed
       
   618 
       
   619 lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)"
       
   620   by (simp add: bl_to_bin_def trunc_bl2bin_aux)
       
   621 
       
   622 lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl"
       
   623   by (simp add: trunc_bl2bin)
       
   624 
       
   625 lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)"
       
   626   apply (rule trans)
       
   627    prefer 2
       
   628    apply (rule trunc_bl2bin [symmetric])
       
   629   apply (cases "k \<le> length bl")
       
   630    apply auto
       
   631   done
       
   632 
       
   633 lemma take_rest_power_bin: "m \<le> n \<Longrightarrow> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)"
       
   634   apply (rule nth_equalityI)
       
   635    apply simp
       
   636   apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin)
       
   637   done
       
   638 
       
   639 lemma last_bin_last': "size xs > 0 \<Longrightarrow> last xs \<longleftrightarrow> bin_last (bl_to_bin_aux xs w)"
       
   640   by (induct xs arbitrary: w) auto
       
   641 
       
   642 lemma last_bin_last: "size xs > 0 \<Longrightarrow> last xs \<longleftrightarrow> bin_last (bl_to_bin xs)"
       
   643   unfolding bl_to_bin_def by (erule last_bin_last')
       
   644 
       
   645 lemma bin_last_last: "bin_last w \<longleftrightarrow> last (bin_to_bl (Suc n) w)"
       
   646   by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt)
       
   647 
       
   648 lemma drop_bin2bl_aux:
       
   649   "drop m (bin_to_bl_aux n bin bs) =
       
   650     bin_to_bl_aux (n - m) bin (drop (m - n) bs)"
       
   651   apply (induction n arbitrary: m bin bs)
       
   652    apply auto
       
   653   apply (case_tac "m \<le> n")
       
   654    apply (auto simp add: not_le Suc_diff_le)
       
   655   apply (case_tac "m - n")
       
   656    apply auto
       
   657   apply (use Suc_diff_Suc in fastforce)
       
   658   done
       
   659 
       
   660 lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin"
       
   661   by (simp add: bin_to_bl_def drop_bin2bl_aux)
       
   662 
       
   663 lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w"
       
   664   apply (induct m arbitrary: w bs)
       
   665    apply clarsimp
       
   666   apply clarsimp
       
   667   apply (simp add: bin_to_bl_aux_alt)
       
   668   apply (simp add: bin_to_bl_def)
       
   669   apply (simp add: bin_to_bl_aux_alt)
       
   670   done
       
   671 
       
   672 lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)"
       
   673   by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp)
       
   674 
       
   675 lemma bin_split_take: "bin_split n c = (a, b) \<Longrightarrow> bin_to_bl m a = take m (bin_to_bl (m + n) c)"
       
   676   apply (induct n arbitrary: b c)
       
   677    apply clarsimp
       
   678   apply (clarsimp simp: Let_def split: prod.split_asm)
       
   679   apply (simp add: bin_to_bl_def)
       
   680   apply (simp add: take_bin2bl_lem drop_bit_Suc)
       
   681   done
       
   682 
       
   683 lemma bin_to_bl_drop_bit:
       
   684   "k = m + n \<Longrightarrow> bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)"
       
   685   using bin_split_take by simp
       
   686 
       
   687 lemma bin_split_take1:
       
   688   "k = m + n \<Longrightarrow> bin_split n c = (a, b) \<Longrightarrow> bin_to_bl m a = take m (bin_to_bl k c)"
       
   689   using bin_split_take by simp
       
   690 
       
   691 lemma bl_bin_bl_rep_drop:
       
   692   "bin_to_bl n (bl_to_bin bl) =
       
   693     replicate (n - length bl) False @ drop (length bl - n) bl"
       
   694   by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin)
       
   695 
       
   696 lemma bl_to_bin_aux_cat:
       
   697   "bl_to_bin_aux bs (bin_cat w nv v) =
       
   698     bin_cat w (nv + length bs) (bl_to_bin_aux bs v)"
       
   699   by (rule bit_eqI)
       
   700     (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps)
       
   701 
       
   702 lemma bin_to_bl_aux_cat:
       
   703   "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs =
       
   704     bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)"
       
   705   by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc)
       
   706 
       
   707 lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)"
       
   708   using bl_to_bin_aux_cat [where nv = "0" and v = "0"]
       
   709   by (simp add: bl_to_bin_def [symmetric])
       
   710 
       
   711 lemma bin_to_bl_cat:
       
   712   "bin_to_bl (nv + nw) (bin_cat v nw w) =
       
   713     bin_to_bl_aux nv v (bin_to_bl nw w)"
       
   714   by (simp add: bin_to_bl_def bin_to_bl_aux_cat)
       
   715 
       
   716 lemmas bl_to_bin_aux_app_cat =
       
   717   trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt]
       
   718 
       
   719 lemmas bin_to_bl_aux_cat_app =
       
   720   trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt]
       
   721 
       
   722 lemma bl_to_bin_app_cat:
       
   723   "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)"
       
   724   by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def)
       
   725 
       
   726 lemma bin_to_bl_cat_app:
       
   727   "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa"
       
   728   by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app)
       
   729 
       
   730 text \<open>\<open>bl_to_bin_app_cat_alt\<close> and \<open>bl_to_bin_app_cat\<close> are easily interderivable.\<close>
       
   731 lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)"
       
   732   by (simp add: bl_to_bin_app_cat)
       
   733 
       
   734 lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1"
       
   735   apply (unfold bl_to_bin_def)
       
   736   apply (induct n)
       
   737    apply simp
       
   738   apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append)
       
   739   apply simp
       
   740   done
       
   741 
       
   742 lemma bin_exhaust:
       
   743   "(\<And>x b. bin = of_bool b + 2 * x \<Longrightarrow> Q) \<Longrightarrow> Q" for bin :: int
       
   744   apply (cases \<open>even bin\<close>)
       
   745    apply (auto elim!: evenE oddE)
       
   746    apply fastforce
       
   747   apply fastforce
       
   748   done
       
   749 
       
   750 primrec rbl_succ :: "bool list \<Rightarrow> bool list"
       
   751   where
       
   752     Nil: "rbl_succ Nil = Nil"
       
   753   | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)"
       
   754 
       
   755 primrec rbl_pred :: "bool list \<Rightarrow> bool list"
       
   756   where
       
   757     Nil: "rbl_pred Nil = Nil"
       
   758   | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)"
       
   759 
       
   760 primrec rbl_add :: "bool list \<Rightarrow> bool list \<Rightarrow> bool list"
       
   761   where \<comment> \<open>result is length of first arg, second arg may be longer\<close>
       
   762     Nil: "rbl_add Nil x = Nil"
       
   763   | Cons: "rbl_add (y # ys) x =
       
   764       (let ws = rbl_add ys (tl x)
       
   765        in (y \<noteq> hd x) # (if hd x \<and> y then rbl_succ ws else ws))"
       
   766 
       
   767 primrec rbl_mult :: "bool list \<Rightarrow> bool list \<Rightarrow> bool list"
       
   768   where \<comment> \<open>result is length of first arg, second arg may be longer\<close>
       
   769     Nil: "rbl_mult Nil x = Nil"
       
   770   | Cons: "rbl_mult (y # ys) x =
       
   771       (let ws = False # rbl_mult ys x
       
   772        in if y then rbl_add ws x else ws)"
       
   773 
       
   774 lemma size_rbl_pred: "length (rbl_pred bl) = length bl"
       
   775   by (induct bl) auto
       
   776 
       
   777 lemma size_rbl_succ: "length (rbl_succ bl) = length bl"
       
   778   by (induct bl) auto
       
   779 
       
   780 lemma size_rbl_add: "length (rbl_add bl cl) = length bl"
       
   781   by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ)
       
   782 
       
   783 lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl"
       
   784   by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add)
       
   785 
       
   786 lemmas rbl_sizes [simp] =
       
   787   size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult
       
   788 
       
   789 lemmas rbl_Nils =
       
   790   rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil
       
   791 
       
   792 lemma rbl_add_app2: "length blb \<ge> length bla \<Longrightarrow> rbl_add bla (blb @ blc) = rbl_add bla blb"
       
   793   apply (induct bla arbitrary: blb)
       
   794    apply simp
       
   795   apply clarsimp
       
   796   apply (case_tac blb, clarsimp)
       
   797   apply (clarsimp simp: Let_def)
       
   798   done
       
   799 
       
   800 lemma rbl_add_take2:
       
   801   "length blb \<ge> length bla \<Longrightarrow> rbl_add bla (take (length bla) blb) = rbl_add bla blb"
       
   802   apply (induct bla arbitrary: blb)
       
   803    apply simp
       
   804   apply clarsimp
       
   805   apply (case_tac blb, clarsimp)
       
   806   apply (clarsimp simp: Let_def)
       
   807   done
       
   808 
       
   809 lemma rbl_mult_app2: "length blb \<ge> length bla \<Longrightarrow> rbl_mult bla (blb @ blc) = rbl_mult bla blb"
       
   810   apply (induct bla arbitrary: blb)
       
   811    apply simp
       
   812   apply clarsimp
       
   813   apply (case_tac blb, clarsimp)
       
   814   apply (clarsimp simp: Let_def rbl_add_app2)
       
   815   done
       
   816 
       
   817 lemma rbl_mult_take2:
       
   818   "length blb \<ge> length bla \<Longrightarrow> rbl_mult bla (take (length bla) blb) = rbl_mult bla blb"
       
   819   apply (rule trans)
       
   820    apply (rule rbl_mult_app2 [symmetric])
       
   821    apply simp
       
   822   apply (rule_tac f = "rbl_mult bla" in arg_cong)
       
   823   apply (rule append_take_drop_id)
       
   824   done
       
   825 
       
   826 lemma rbl_add_split:
       
   827   "P (rbl_add (y # ys) (x # xs)) =
       
   828     (\<forall>ws. length ws = length ys \<longrightarrow> ws = rbl_add ys xs \<longrightarrow>
       
   829       (y \<longrightarrow> ((x \<longrightarrow> P (False # rbl_succ ws)) \<and> (\<not> x \<longrightarrow> P (True # ws)))) \<and>
       
   830       (\<not> y \<longrightarrow> P (x # ws)))"
       
   831   by (cases y) (auto simp: Let_def)
       
   832 
       
   833 lemma rbl_mult_split:
       
   834   "P (rbl_mult (y # ys) xs) =
       
   835     (\<forall>ws. length ws = Suc (length ys) \<longrightarrow> ws = False # rbl_mult ys xs \<longrightarrow>
       
   836       (y \<longrightarrow> P (rbl_add ws xs)) \<and> (\<not> y \<longrightarrow> P ws))"
       
   837   by (auto simp: Let_def)
       
   838 
       
   839 lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))"
       
   840 proof (unfold bin_to_bl_def, induction n arbitrary: bin)
       
   841   case 0
       
   842   then show ?case
       
   843     by simp
       
   844 next
       
   845   case (Suc n)
       
   846   obtain b k where \<open>bin = of_bool b + 2 * k\<close>
       
   847     using bin_exhaust by blast
       
   848   moreover have \<open>(2 * k - 1) div 2 = k - 1\<close>
       
   849     using even_succ_div_2 [of \<open>2 * (k - 1)\<close>] 
       
   850     by simp
       
   851   ultimately show ?case
       
   852     using Suc [of \<open>bin div 2\<close>]
       
   853     by simp (simp add: bin_to_bl_aux_alt)
       
   854 qed
       
   855 
       
   856 lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))"
       
   857   apply (unfold bin_to_bl_def)
       
   858   apply (induction n arbitrary: bin)
       
   859    apply simp_all
       
   860   apply (case_tac bin rule: bin_exhaust)
       
   861   apply simp
       
   862   apply (simp add: bin_to_bl_aux_alt ac_simps)
       
   863   done
       
   864 
       
   865 lemma rbl_add:
       
   866   "\<And>bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) =
       
   867     rev (bin_to_bl n (bina + binb))"
       
   868   apply (unfold bin_to_bl_def)
       
   869   apply (induct n)
       
   870    apply simp
       
   871   apply clarsimp
       
   872   apply (case_tac bina rule: bin_exhaust)
       
   873   apply (case_tac binb rule: bin_exhaust)
       
   874   apply (case_tac b)
       
   875    apply (case_tac [!] "ba")
       
   876      apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps)
       
   877   done
       
   878 
       
   879 lemma rbl_add_long:
       
   880   "m \<ge> n \<Longrightarrow> rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) =
       
   881     rev (bin_to_bl n (bina + binb))"
       
   882   apply (rule box_equals [OF _ rbl_add_take2 rbl_add])
       
   883    apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong)
       
   884    apply (rule rev_swap [THEN iffD1])
       
   885    apply (simp add: rev_take drop_bin2bl)
       
   886   apply simp
       
   887   done
       
   888 
       
   889 lemma rbl_mult_gt1:
       
   890   "m \<ge> length bl \<Longrightarrow>
       
   891     rbl_mult bl (rev (bin_to_bl m binb)) =
       
   892     rbl_mult bl (rev (bin_to_bl (length bl) binb))"
       
   893   apply (rule trans)
       
   894    apply (rule rbl_mult_take2 [symmetric])
       
   895    apply simp_all
       
   896   apply (rule_tac f = "rbl_mult bl" in arg_cong)
       
   897   apply (rule rev_swap [THEN iffD1])
       
   898   apply (simp add: rev_take drop_bin2bl)
       
   899   done
       
   900 
       
   901 lemma rbl_mult_gt:
       
   902   "m > n \<Longrightarrow>
       
   903     rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) =
       
   904     rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))"
       
   905   by (auto intro: trans [OF rbl_mult_gt1])
       
   906 
       
   907 lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt]
       
   908 
       
   909 lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))"
       
   910   by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt)
       
   911 
       
   912 lemma rbl_mult:
       
   913   "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) =
       
   914     rev (bin_to_bl n (bina * binb))"
       
   915   apply (induct n arbitrary: bina binb)
       
   916    apply simp_all
       
   917   apply (unfold bin_to_bl_def)
       
   918   apply clarsimp
       
   919   apply (case_tac bina rule: bin_exhaust)
       
   920   apply (case_tac binb rule: bin_exhaust)
       
   921   apply simp
       
   922   apply (simp add: bin_to_bl_aux_alt)
       
   923   apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps)
       
   924   done
       
   925 
       
   926 lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n"
       
   927   by (induct xs) auto
       
   928 
       
   929 lemma bin_cat_foldl_lem:
       
   930   "foldl (\<lambda>u. bin_cat u n) x xs =
       
   931     bin_cat x (size xs * n) (foldl (\<lambda>u. bin_cat u n) y xs)"
       
   932   apply (induct xs arbitrary: x)
       
   933    apply simp
       
   934   apply (simp (no_asm))
       
   935   apply (frule asm_rl)
       
   936   apply (drule meta_spec)
       
   937   apply (erule trans)
       
   938   apply (drule_tac x = "bin_cat y n a" in meta_spec)
       
   939   apply (simp add: bin_cat_assoc_sym min.absorb2)
       
   940   done
       
   941 
       
   942 lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))"
       
   943   apply (unfold bin_rcat_eq_foldl)
       
   944   apply (rule sym)
       
   945   apply (induct wl)
       
   946    apply (auto simp add: bl_to_bin_append)
       
   947   apply (simp add: bl_to_bin_aux_alt sclem)
       
   948   apply (simp add: bin_cat_foldl_lem [symmetric])
       
   949   done
       
   950 
       
   951 lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \<longleftrightarrow> bs \<noteq> [] \<and> last bs"
       
   952 by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0])
       
   953 
       
   954 lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)"
       
   955 by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux)
       
   956 
       
   957 lemma bl_xor_aux_bin:
       
   958   "map2 (\<lambda>x y. x \<noteq> y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) =
       
   959     bin_to_bl_aux n (v XOR w) (map2 (\<lambda>x y. x \<noteq> y) bs cs)"
       
   960   apply (induction n arbitrary: v w bs cs)
       
   961    apply auto
       
   962   apply (case_tac v rule: bin_exhaust)
       
   963   apply (case_tac w rule: bin_exhaust)
       
   964   apply clarsimp
       
   965   done
       
   966 
       
   967 lemma bl_or_aux_bin:
       
   968   "map2 (\<or>) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) =
       
   969     bin_to_bl_aux n (v OR w) (map2 (\<or>) bs cs)"
       
   970   by (induct n arbitrary: v w bs cs) simp_all
       
   971 
       
   972 lemma bl_and_aux_bin:
       
   973   "map2 (\<and>) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) =
       
   974     bin_to_bl_aux n (v AND w) (map2 (\<and>) bs cs)"
       
   975   by (induction n arbitrary: v w bs cs) simp_all
       
   976 
       
   977 lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)"
       
   978   by (induct n arbitrary: w cs) auto
       
   979 
       
   980 lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)"
       
   981   by (simp add: bin_to_bl_def bl_not_aux_bin)
       
   982 
       
   983 lemma bl_and_bin: "map2 (\<and>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)"
       
   984   by (simp add: bin_to_bl_def bl_and_aux_bin)
       
   985 
       
   986 lemma bl_or_bin: "map2 (\<or>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)"
       
   987   by (simp add: bin_to_bl_def bl_or_aux_bin)
       
   988 
       
   989 lemma bl_xor_bin: "map2 (\<noteq>) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)"
       
   990   using bl_xor_aux_bin by (simp add: bin_to_bl_def)
       
   991 
       
   992 
       
   993 subsection \<open>Type \<^typ>\<open>'a word\<close>\<close>
       
   994 
       
   995 lift_definition of_bl :: \<open>bool list \<Rightarrow> 'a::len word\<close>
       
   996   is bl_to_bin .
       
   997 
       
   998 lift_definition to_bl :: \<open>'a::len word \<Rightarrow> bool list\<close>
       
   999   is \<open>bin_to_bl LENGTH('a)\<close>
       
  1000   by (simp add: bl_to_bin_inj)
       
  1001 
       
  1002 lemma to_bl_eq:
       
  1003   \<open>to_bl w = bin_to_bl (LENGTH('a)) (uint w)\<close>
       
  1004   for w :: \<open>'a::len word\<close>
       
  1005   by transfer simp
       
  1006 
       
  1007 lemma bit_of_bl_iff:
       
  1008   \<open>bit (of_bl bs :: 'a word) n \<longleftrightarrow> rev bs ! n \<and> n < LENGTH('a::len) \<and> n < length bs\<close>
       
  1009   by transfer (simp add: bin_nth_of_bl ac_simps)
       
  1010 
       
  1011 lemma rev_to_bl_eq:
       
  1012   \<open>rev (to_bl w) = map (bit w) [0..<LENGTH('a)]\<close>
       
  1013   for w :: \<open>'a::len word\<close>
       
  1014   apply (rule nth_equalityI)
       
  1015    apply (simp add: to_bl.rep_eq)
       
  1016   apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq)
       
  1017   done
       
  1018 
       
  1019 lemma to_bl_eq_rev:
       
  1020   \<open>to_bl w = map (bit w) (rev [0..<LENGTH('a)])\<close>
       
  1021   for w :: \<open>'a::len word\<close>
       
  1022   using rev_to_bl_eq [of w]
       
  1023   apply (subst rev_is_rev_conv [symmetric])
       
  1024   apply (simp add: rev_map)
       
  1025   done
       
  1026 
       
  1027 lemma of_bl_rev_eq:
       
  1028   \<open>of_bl (rev bs) = horner_sum of_bool 2 bs\<close>
       
  1029   apply (rule bit_word_eqI)
       
  1030   apply (simp add: bit_of_bl_iff)
       
  1031   apply transfer
       
  1032   apply (simp add: bit_horner_sum_bit_iff ac_simps)
       
  1033   done
       
  1034 
       
  1035 lemma of_bl_eq:
       
  1036   \<open>of_bl bs = horner_sum of_bool 2 (rev bs)\<close>
       
  1037   using of_bl_rev_eq [of \<open>rev bs\<close>] by simp
       
  1038 
       
  1039 lemma bshiftr1_eq:
       
  1040   \<open>bshiftr1 b w = of_bl (b # butlast (to_bl w))\<close>
       
  1041   apply transfer
       
  1042   apply auto
       
  1043    apply (subst bl_to_bin_app_cat [of \<open>[True]\<close>, simplified])
       
  1044    apply simp
       
  1045    apply (metis One_nat_def add.commute bin_bl_bin bin_last_bl_to_bin bin_rest_bl_to_bin butlast_bin_rest concat_bit_eq last.simps list.distinct(1) list.size(3) list.size(4) odd_iff_mod_2_eq_one plus_1_eq_Suc power_Suc0_right push_bit_of_1 size_bin_to_bl take_bit_eq_mod trunc_bl2bin_len)
       
  1046   apply (simp add: butlast_rest_bl2bin)
       
  1047   done
       
  1048 
       
  1049 lemma length_to_bl_eq:
       
  1050   \<open>length (to_bl w) = LENGTH('a)\<close>
       
  1051   for w :: \<open>'a::len word\<close>
       
  1052   by transfer simp
       
  1053 
       
  1054 lemma word_rotr_eq:
       
  1055   \<open>word_rotr n w = of_bl (rotater n (to_bl w))\<close>
       
  1056   apply (rule bit_word_eqI)
       
  1057   subgoal for n
       
  1058     apply (cases \<open>n < LENGTH('a)\<close>)
       
  1059      apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps)
       
  1060     done
       
  1061   done
       
  1062 
       
  1063 lemma word_rotl_eq:
       
  1064   \<open>word_rotl n w = of_bl (rotate n (to_bl w))\<close>
       
  1065 proof -
       
  1066   have \<open>rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\<close>
       
  1067     by (simp add: rotater_rev')
       
  1068   then show ?thesis
       
  1069     apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq)
       
  1070     apply (rule bit_word_eqI)
       
  1071     subgoal for n
       
  1072       apply (cases \<open>n < LENGTH('a)\<close>)
       
  1073        apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater)
       
  1074       done
       
  1075     done
       
  1076 qed
       
  1077 
       
  1078 lemma to_bl_def': "(to_bl :: 'a::len word \<Rightarrow> bool list) = bin_to_bl (LENGTH('a)) \<circ> uint"
       
  1079   by transfer (simp add: fun_eq_iff)
       
  1080 
       
  1081 \<comment> \<open>type definitions theorem for in terms of equivalent bool list\<close>
       
  1082 lemma td_bl:
       
  1083   "type_definition
       
  1084     (to_bl :: 'a::len word \<Rightarrow> bool list)
       
  1085     of_bl
       
  1086     {bl. length bl = LENGTH('a)}"
       
  1087   apply (standard; transfer)
       
  1088   apply (auto dest: sym)
       
  1089   done
       
  1090 
       
  1091 interpretation word_bl:
       
  1092   type_definition
       
  1093     "to_bl :: 'a::len word \<Rightarrow> bool list"
       
  1094     of_bl
       
  1095     "{bl. length bl = LENGTH('a::len)}"
       
  1096   by (fact td_bl)
       
  1097 
       
  1098 lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff]
       
  1099 
       
  1100 lemma word_size_bl: "size w = size (to_bl w)"
       
  1101   by (auto simp: word_size)
       
  1102 
       
  1103 lemma to_bl_use_of_bl: "to_bl w = bl \<longleftrightarrow> w = of_bl bl \<and> length bl = length (to_bl w)"
       
  1104   by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq])
       
  1105 
       
  1106 lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)"
       
  1107   for x :: "'a::len word"
       
  1108   unfolding word_bl_Rep' by (rule len_gt_0)
       
  1109 
       
  1110 lemma bl_not_Nil [iff]: "to_bl x \<noteq> []"
       
  1111   for x :: "'a::len word"
       
  1112   by (fact length_bl_gt_0 [unfolded length_greater_0_conv])
       
  1113 
       
  1114 lemma length_bl_neq_0 [iff]: "length (to_bl x) \<noteq> 0"
       
  1115   for x :: "'a::len word"
       
  1116   by (fact length_bl_gt_0 [THEN gr_implies_not0])
       
  1117 
       
  1118 lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)"
       
  1119   apply transfer
       
  1120   apply (auto simp add: bin_sign_def)
       
  1121   using bin_sign_lem bl_sbin_sign apply fastforce
       
  1122   using bin_sign_lem bl_sbin_sign apply force
       
  1123   done
       
  1124 
       
  1125 lemma of_bl_drop':
       
  1126   "lend = length bl - LENGTH('a::len) \<Longrightarrow>
       
  1127     of_bl (drop lend bl) = (of_bl bl :: 'a word)"
       
  1128   by transfer (simp flip: trunc_bl2bin)
       
  1129 
       
  1130 lemma test_bit_of_bl:
       
  1131   "(of_bl bl::'a::len word) !! n = (rev bl ! n \<and> n < LENGTH('a) \<and> n < length bl)"
       
  1132   by transfer (simp add: bin_nth_of_bl ac_simps)
       
  1133 
       
  1134 lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))"
       
  1135   by transfer simp
       
  1136 
       
  1137 lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)"
       
  1138   by transfer simp
       
  1139 
       
  1140 lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w"
       
  1141   by (simp add: uint_bl word_size)
       
  1142 
       
  1143 lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin"
       
  1144   by (auto simp: uint_bl word_ubin.eq_norm word_size)
       
  1145 
       
  1146 lemma to_bl_numeral [simp]:
       
  1147   "to_bl (numeral bin::'a::len word) =
       
  1148     bin_to_bl (LENGTH('a)) (numeral bin)"
       
  1149   unfolding word_numeral_alt by (rule to_bl_of_bin)
       
  1150 
       
  1151 lemma to_bl_neg_numeral [simp]:
       
  1152   "to_bl (- numeral bin::'a::len word) =
       
  1153     bin_to_bl (LENGTH('a)) (- numeral bin)"
       
  1154   unfolding word_neg_numeral_alt by (rule to_bl_of_bin)
       
  1155 
       
  1156 lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w"
       
  1157   by (simp add: uint_bl word_size)
       
  1158 
       
  1159 lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x"
       
  1160   for x :: "'a::len word"
       
  1161   by (rule trans [OF bin_bl_bin word_ubin.norm_Rep])
       
  1162 
       
  1163 lemma ucast_bl: "ucast w = of_bl (to_bl w)"
       
  1164   by transfer simp
       
  1165 
       
  1166 lemma ucast_down_bl:
       
  1167   \<open>(ucast :: 'a::len word \<Rightarrow> 'b::len word) (of_bl bl) = of_bl bl\<close>
       
  1168     if \<open>is_down (ucast :: 'a::len word \<Rightarrow> 'b::len word)\<close>
       
  1169   using that by transfer simp
       
  1170 
       
  1171 lemma of_bl_append_same: "of_bl (X @ to_bl w) = w"
       
  1172   by transfer (simp add: bl_to_bin_app_cat) 
       
  1173 
       
  1174 lemma ucast_of_bl_up:
       
  1175   \<open>ucast (of_bl bl :: 'a::len word) = of_bl bl\<close>
       
  1176   if \<open>size bl \<le> size (of_bl bl :: 'a::len word)\<close>
       
  1177   using that
       
  1178   apply transfer
       
  1179   apply (rule bit_eqI)
       
  1180   apply (auto simp add: bit_take_bit_iff)
       
  1181   apply (subst (asm) trunc_bl2bin_len [symmetric])
       
  1182   apply (auto simp only: bit_take_bit_iff)
       
  1183   done
       
  1184 
       
  1185 lemma word_rev_tf:
       
  1186   "to_bl (of_bl bl::'a::len word) =
       
  1187     rev (takefill False (LENGTH('a)) (rev bl))"
       
  1188   by transfer (simp add: bl_bin_bl_rtf)
       
  1189 
       
  1190 lemma word_rep_drop:
       
  1191   "to_bl (of_bl bl::'a::len word) =
       
  1192     replicate (LENGTH('a) - length bl) False @
       
  1193     drop (length bl - LENGTH('a)) bl"
       
  1194   by (simp add: word_rev_tf takefill_alt rev_take)
       
  1195 
       
  1196 lemma to_bl_ucast:
       
  1197   "to_bl (ucast (w::'b::len word) ::'a::len word) =
       
  1198     replicate (LENGTH('a) - LENGTH('b)) False @
       
  1199     drop (LENGTH('b) - LENGTH('a)) (to_bl w)"
       
  1200   apply (unfold ucast_bl)
       
  1201   apply (rule trans)
       
  1202    apply (rule word_rep_drop)
       
  1203   apply simp
       
  1204   done
       
  1205 
       
  1206 lemma ucast_up_app:
       
  1207   \<open>to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\<close>
       
  1208     if \<open>source_size (ucast :: 'a word \<Rightarrow> 'b word) + n = target_size (ucast :: 'a word \<Rightarrow> 'b word)\<close>
       
  1209     for w :: \<open>'a::len word\<close>
       
  1210   using that
       
  1211   by (auto simp add : source_size target_size to_bl_ucast)
       
  1212 
       
  1213 lemma ucast_down_drop [OF refl]:
       
  1214   "uc = ucast \<Longrightarrow> source_size uc = target_size uc + n \<Longrightarrow>
       
  1215     to_bl (uc w) = drop n (to_bl w)"
       
  1216   by (auto simp add : source_size target_size to_bl_ucast)
       
  1217 
       
  1218 lemma scast_down_drop [OF refl]:
       
  1219   "sc = scast \<Longrightarrow> source_size sc = target_size sc + n \<Longrightarrow>
       
  1220     to_bl (sc w) = drop n (to_bl w)"
       
  1221   apply (subgoal_tac "sc = ucast")
       
  1222    apply safe
       
  1223    apply simp
       
  1224    apply (erule ucast_down_drop)
       
  1225   apply (rule down_cast_same [symmetric])
       
  1226   apply (simp add : source_size target_size is_down)
       
  1227   done
       
  1228 
       
  1229 lemma word_0_bl [simp]: "of_bl [] = 0"
       
  1230   by transfer simp
       
  1231 
       
  1232 lemma word_1_bl: "of_bl [True] = 1"
       
  1233   by transfer (simp add: bl_to_bin_def)
       
  1234 
       
  1235 lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0"
       
  1236   by transfer (simp add: bl_to_bin_rep_False)
       
  1237 
       
  1238 lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False"
       
  1239   by (simp add: uint_bl word_size bin_to_bl_zero)
       
  1240 
       
  1241 \<comment> \<open>links with \<open>rbl\<close> operations\<close>
       
  1242 lemma word_succ_rbl: "to_bl w = bl \<Longrightarrow> to_bl (word_succ w) = rev (rbl_succ (rev bl))"
       
  1243   by transfer (simp add: rbl_succ)
       
  1244 
       
  1245 lemma word_pred_rbl: "to_bl w = bl \<Longrightarrow> to_bl (word_pred w) = rev (rbl_pred (rev bl))"
       
  1246   by transfer (simp add: rbl_pred)
       
  1247 
       
  1248 lemma word_add_rbl:
       
  1249   "to_bl v = vbl \<Longrightarrow> to_bl w = wbl \<Longrightarrow>
       
  1250     to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))"
       
  1251   apply transfer
       
  1252   apply (drule sym)
       
  1253   apply (drule sym)
       
  1254   apply (simp add: rbl_add)
       
  1255   done
       
  1256 
       
  1257 lemma word_mult_rbl:
       
  1258   "to_bl v = vbl \<Longrightarrow> to_bl w = wbl \<Longrightarrow>
       
  1259     to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))"
       
  1260   apply transfer
       
  1261   apply (drule sym)
       
  1262   apply (drule sym)
       
  1263   apply (simp add: rbl_mult)
       
  1264   done
       
  1265 
       
  1266 lemma rtb_rbl_ariths:
       
  1267   "rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_succ w)) = rbl_succ ys"
       
  1268   "rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_pred w)) = rbl_pred ys"
       
  1269   "rev (to_bl v) = ys \<Longrightarrow> rev (to_bl w) = xs \<Longrightarrow> rev (to_bl (v * w)) = rbl_mult ys xs"
       
  1270   "rev (to_bl v) = ys \<Longrightarrow> rev (to_bl w) = xs \<Longrightarrow> rev (to_bl (v + w)) = rbl_add ys xs"
       
  1271   by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl)
       
  1272 
       
  1273 lemma of_bl_length_less:
       
  1274   \<open>(of_bl x :: 'a::len word) < 2 ^ k\<close>
       
  1275     if \<open>length x = k\<close> \<open>k < LENGTH('a)\<close>
       
  1276 proof -
       
  1277   from that have \<open>length x < LENGTH('a)\<close>
       
  1278     by simp
       
  1279   then have \<open>(of_bl x :: 'a::len word) < 2 ^ length x\<close>
       
  1280     apply (simp add: of_bl_eq)
       
  1281     apply transfer
       
  1282     apply (simp add: take_bit_horner_sum_bit_eq)
       
  1283     apply (subst length_rev [symmetric])
       
  1284     apply (simp only: horner_sum_of_bool_2_less)
       
  1285     done
       
  1286   with that show ?thesis
       
  1287     by simp
       
  1288 qed
       
  1289 
       
  1290 lemma word_eq_rbl_eq: "x = y \<longleftrightarrow> rev (to_bl x) = rev (to_bl y)"
       
  1291   by simp
       
  1292 
       
  1293 lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)"
       
  1294   by transfer (simp add: bl_not_bin)
       
  1295 
       
  1296 lemma bl_word_xor: "to_bl (v XOR w) = map2 (\<noteq>) (to_bl v) (to_bl w)"
       
  1297   by transfer (simp flip: bl_xor_bin)
       
  1298 
       
  1299 lemma bl_word_or: "to_bl (v OR w) = map2 (\<or>) (to_bl v) (to_bl w)"
       
  1300   by transfer (simp flip: bl_or_bin)
       
  1301 
       
  1302 lemma bl_word_and: "to_bl (v AND w) = map2 (\<and>) (to_bl v) (to_bl w)"
       
  1303   by transfer (simp flip: bl_and_bin)
       
  1304 
       
  1305 lemma bin_nth_uint': "bin_nth (uint w) n \<longleftrightarrow> rev (bin_to_bl (size w) (uint w)) ! n \<and> n < size w"
       
  1306   apply (unfold word_size)
       
  1307   apply (safe elim!: bin_nth_uint_imp)
       
  1308    apply (frule bin_nth_uint_imp)
       
  1309    apply (fast dest!: bin_nth_bl)+
       
  1310   done
       
  1311 
       
  1312 lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size]
       
  1313 
       
  1314 lemma test_bit_bl: "w !! n \<longleftrightarrow> rev (to_bl w) ! n \<and> n < size w"
       
  1315   by transfer (auto simp add: bin_nth_bl)
       
  1316 
       
  1317 lemma to_bl_nth: "n < size w \<Longrightarrow> to_bl w ! n = w !! (size w - Suc n)"
       
  1318   by (simp add: word_size rev_nth test_bit_bl)
       
  1319 
       
  1320 lemma map_bit_interval_eq:
       
  1321   \<open>map (bit w) [0..<n] = takefill False n (rev (to_bl w))\<close> for w :: \<open>'a::len word\<close>
       
  1322 proof (rule nth_equalityI)
       
  1323   show \<open>length (map (bit w) [0..<n]) =
       
  1324     length (takefill False n (rev (to_bl w)))\<close>
       
  1325     by simp
       
  1326   fix m
       
  1327   assume \<open>m < length (map (bit w) [0..<n])\<close>
       
  1328   then have \<open>m < n\<close>
       
  1329     by simp
       
  1330   then have \<open>bit w m \<longleftrightarrow> takefill False n (rev (to_bl w)) ! m\<close>
       
  1331     by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length)
       
  1332   with \<open>m < n \<close>show \<open>map (bit w) [0..<n] ! m \<longleftrightarrow> takefill False n (rev (to_bl w)) ! m\<close>
       
  1333     by simp
       
  1334 qed
       
  1335 
       
  1336 lemma to_bl_unfold:
       
  1337   \<open>to_bl w = rev (map (bit w) [0..<LENGTH('a)])\<close> for w :: \<open>'a::len word\<close>
       
  1338   by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def)
       
  1339 
       
  1340 lemma nth_rev_to_bl:
       
  1341   \<open>rev (to_bl w) ! n \<longleftrightarrow> bit w n\<close>
       
  1342   if \<open>n < LENGTH('a)\<close> for w :: \<open>'a::len word\<close>
       
  1343   using that by (simp add: to_bl_unfold)
       
  1344 
       
  1345 lemma nth_to_bl:
       
  1346   \<open>to_bl w ! n \<longleftrightarrow> bit w (LENGTH('a) - Suc n)\<close>
       
  1347   if \<open>n < LENGTH('a)\<close> for w :: \<open>'a::len word\<close>
       
  1348   using that by (simp add: to_bl_unfold rev_nth)
       
  1349 
       
  1350 lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs"
       
  1351   by (auto simp: of_bl_def bl_to_bin_rep_F)
       
  1352 
       
  1353 lemma [code abstract]:
       
  1354   \<open>uint (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\<close>
       
  1355   apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq)
       
  1356   apply transfer
       
  1357   apply simp
       
  1358   done
       
  1359 
       
  1360 lemma [code]:
       
  1361   \<open>to_bl w = map (bit w) (rev [0..<LENGTH('a::len)])\<close>
       
  1362   for w :: \<open>'a::len word\<close>
       
  1363   by (simp add: to_bl_unfold rev_map)
       
  1364 
       
  1365 lemma word_reverse_eq_of_bl_rev_to_bl:
       
  1366   \<open>word_reverse w = of_bl (rev (to_bl w))\<close>
       
  1367   by (rule bit_word_eqI)
       
  1368     (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl)
       
  1369 
       
  1370 lemmas word_reverse_no_def [simp] =
       
  1371   word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w
       
  1372 
       
  1373 lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)"
       
  1374   by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq)
       
  1375 
       
  1376 lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True"
       
  1377   apply (rule word_bl.Abs_inverse')
       
  1378    apply simp
       
  1379   apply (rule word_eqI)
       
  1380   apply (clarsimp simp add: word_size)
       
  1381   apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size)
       
  1382   done
       
  1383 
       
  1384 lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\<or>) (rev (to_bl x)) (rev (to_bl y))"
       
  1385   by (simp add: zip_rev bl_word_or rev_map)
       
  1386 
       
  1387 lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\<and>) (rev (to_bl x)) (rev (to_bl y))"
       
  1388   by (simp add: zip_rev bl_word_and rev_map)
       
  1389 
       
  1390 lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\<noteq>) (rev (to_bl x)) (rev (to_bl y))"
       
  1391   by (simp add: zip_rev bl_word_xor rev_map)
       
  1392 
       
  1393 lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))"
       
  1394   by (simp add: bl_word_not rev_map)
       
  1395 
       
  1396 lemma bshiftr1_numeral [simp]:
       
  1397   \<open>bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\<close>
       
  1398   by (simp add: bshiftr1_eq)
       
  1399 
       
  1400 lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)"
       
  1401   unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp
       
  1402 
       
  1403 lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])"
       
  1404   by transfer (simp add: bl_to_bin_append)
       
  1405 
       
  1406 lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])"
       
  1407   for w :: "'a::len word"
       
  1408 proof -
       
  1409   have "shiftl1 w = shiftl1 (of_bl (to_bl w))"
       
  1410     by simp
       
  1411   also have "\<dots> = of_bl (to_bl w @ [False])"
       
  1412     by (rule shiftl1_of_bl)
       
  1413   finally show ?thesis .
       
  1414 qed
       
  1415 
       
  1416 lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]"
       
  1417   for w :: "'a::len word"
       
  1418   by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI)
       
  1419 
       
  1420 \<comment> \<open>Generalized version of \<open>bl_shiftl1\<close>. Maybe this one should replace it?\<close>
       
  1421 lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])"
       
  1422   by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append)
       
  1423 
       
  1424 lemma shiftr1_bl:
       
  1425   \<open>shiftr1 w = of_bl (butlast (to_bl w))\<close>
       
  1426 proof (rule bit_word_eqI)
       
  1427   fix n
       
  1428   assume \<open>n < LENGTH('a)\<close>
       
  1429   show \<open>bit (shiftr1 w) n \<longleftrightarrow> bit (of_bl (butlast (to_bl w)) :: 'a word) n\<close>
       
  1430   proof (cases \<open>n = LENGTH('a) - 1\<close>)
       
  1431     case True
       
  1432     then show ?thesis
       
  1433       by (simp add: bit_shiftr1_iff bit_of_bl_iff)
       
  1434   next
       
  1435     case False
       
  1436     with \<open>n < LENGTH('a)\<close>
       
  1437     have \<open>n < LENGTH('a) - 1\<close>
       
  1438       by simp
       
  1439     with \<open>n < LENGTH('a)\<close> show ?thesis 
       
  1440       by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast
       
  1441         word_size test_bit_word_eq to_bl_nth)
       
  1442   qed
       
  1443 qed
       
  1444 
       
  1445 lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)"
       
  1446   for w :: "'a::len word"
       
  1447   by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI])
       
  1448 
       
  1449 \<comment> \<open>Generalized version of \<open>bl_shiftr1\<close>. Maybe this one should replace it?\<close>
       
  1450 lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)"
       
  1451   apply (rule word_bl.Abs_inverse')
       
  1452    apply (simp del: butlast.simps)
       
  1453   apply (simp add: shiftr1_bl of_bl_def)
       
  1454   done
       
  1455 
       
  1456 lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)"
       
  1457   for w :: "'a::len word"
       
  1458 proof (rule nth_equalityI)
       
  1459   fix n
       
  1460   assume \<open>n < length (to_bl (sshiftr1 w))\<close>
       
  1461   then have \<open>n < LENGTH('a)\<close>
       
  1462     by simp
       
  1463   then show \<open>to_bl (sshiftr1 w) ! n \<longleftrightarrow> (hd (to_bl w) # butlast (to_bl w)) ! n\<close>
       
  1464     apply (cases n)
       
  1465      apply (simp_all add: to_bl_nth word_size hd_conv_nth test_bit_eq_bit bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl)
       
  1466     done
       
  1467 qed simp
       
  1468 
       
  1469 lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)"
       
  1470   for w :: "'a::len word"
       
  1471   apply (unfold shiftr_def)
       
  1472   apply (induct n)
       
  1473    prefer 2
       
  1474    apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric])
       
  1475    apply (rule butlast_take [THEN trans])
       
  1476     apply (auto simp: word_size)
       
  1477   done
       
  1478 
       
  1479 lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)"
       
  1480   for w :: "'a::len word"
       
  1481    apply (simp_all add: word_size sshiftr_eq)
       
  1482   apply (rule nth_equalityI)
       
  1483    apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff)
       
  1484   done
       
  1485 
       
  1486 lemma take_shiftr: "n \<le> size w \<Longrightarrow> take n (to_bl (w >> n)) = replicate n False"
       
  1487   apply (unfold shiftr_def)
       
  1488   apply (induct n)
       
  1489    prefer 2
       
  1490    apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size)
       
  1491    apply (rule take_butlast [THEN trans])
       
  1492     apply (auto simp: word_size)
       
  1493   done
       
  1494 
       
  1495 lemma take_sshiftr':
       
  1496   "n \<le> size w \<Longrightarrow> hd (to_bl (w >>> n)) = hd (to_bl w) \<and>
       
  1497     take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))"
       
  1498   for w :: "'a::len word"
       
  1499   apply (auto simp add: sshiftr_eq hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq)
       
  1500   apply (rule nth_equalityI)
       
  1501     apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff)
       
  1502   apply (rule nth_equalityI)
       
  1503    apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff)
       
  1504   done
       
  1505 
       
  1506 lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1]
       
  1507 lemmas take_sshiftr = take_sshiftr' [THEN conjunct2]
       
  1508 
       
  1509 lemma atd_lem: "take n xs = t \<Longrightarrow> drop n xs = d \<Longrightarrow> xs = t @ d"
       
  1510   by (auto intro: append_take_drop_id [symmetric])
       
  1511 
       
  1512 lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr]
       
  1513 lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr]
       
  1514 
       
  1515 lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)"
       
  1516   by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same)
       
  1517 
       
  1518 lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)"
       
  1519   for w :: "'a::len word"
       
  1520 proof -
       
  1521   have "w << n = of_bl (to_bl w) << n"
       
  1522     by simp
       
  1523   also have "\<dots> = of_bl (to_bl w @ replicate n False)"
       
  1524     by (rule shiftl_of_bl)
       
  1525   finally show ?thesis .
       
  1526 qed
       
  1527 
       
  1528 lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False"
       
  1529   by (simp add: shiftl_bl word_rep_drop word_size)
       
  1530 
       
  1531 lemma shiftr1_bl_of:
       
  1532   "length bl \<le> LENGTH('a) \<Longrightarrow>
       
  1533     shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)"
       
  1534   by transfer (simp add: butlast_rest_bl2bin trunc_bl2bin)
       
  1535 
       
  1536 lemma shiftr_bl_of:
       
  1537   "length bl \<le> LENGTH('a) \<Longrightarrow>
       
  1538     (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)"
       
  1539   apply (unfold shiftr_def)
       
  1540   apply (induct n)
       
  1541    apply clarsimp
       
  1542   apply clarsimp
       
  1543   apply (subst shiftr1_bl_of)
       
  1544    apply simp
       
  1545   apply (simp add: butlast_take)
       
  1546   done
       
  1547 
       
  1548 lemma shiftr_bl: "x >> n \<equiv> of_bl (take (LENGTH('a) - n) (to_bl x))"
       
  1549   for x :: "'a::len word"
       
  1550   using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp
       
  1551 
       
  1552 lemma aligned_bl_add_size [OF refl]:
       
  1553   "size x - n = m \<Longrightarrow> n \<le> size x \<Longrightarrow> drop m (to_bl x) = replicate n False \<Longrightarrow>
       
  1554     take m (to_bl y) = replicate m False \<Longrightarrow>
       
  1555     to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \<open>'a::len word\<close>
       
  1556   apply (subgoal_tac "x AND y = 0")
       
  1557    prefer 2
       
  1558    apply (rule word_bl.Rep_eqD)
       
  1559    apply (simp add: bl_word_and)
       
  1560    apply (rule align_lem_and [THEN trans])
       
  1561        apply (simp_all add: word_size)[5]
       
  1562    apply simp
       
  1563   apply (subst word_plus_and_or [symmetric])
       
  1564   apply (simp add : bl_word_or)
       
  1565   apply (rule align_lem_or)
       
  1566      apply (simp_all add: word_size)
       
  1567   done
       
  1568 
       
  1569 lemma mask_bl: "mask n = of_bl (replicate n True)"
       
  1570   by (auto simp add : test_bit_of_bl word_size intro: word_eqI)
       
  1571 
       
  1572 lemma bl_and_mask':
       
  1573   "to_bl (w AND mask n :: 'a::len word) =
       
  1574     replicate (LENGTH('a) - n) False @
       
  1575     drop (LENGTH('a) - n) (to_bl w)"
       
  1576   apply (rule nth_equalityI)
       
  1577    apply simp
       
  1578   apply (clarsimp simp add: to_bl_nth word_size)
       
  1579   apply (simp add: word_size word_ops_nth_size)
       
  1580   apply (auto simp add: word_size test_bit_bl nth_append rev_nth)
       
  1581   done
       
  1582 
       
  1583 lemma slice1_eq_of_bl:
       
  1584   \<open>(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\<close>
       
  1585   for w :: \<open>'a::len word\<close>
       
  1586 proof (rule bit_word_eqI)
       
  1587   fix m
       
  1588   assume \<open>m < LENGTH('b)\<close>
       
  1589   show \<open>bit (slice1 n w :: 'b::len word) m \<longleftrightarrow> bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\<close>
       
  1590     by (cases \<open>m \<ge> n\<close>; cases \<open>LENGTH('a) \<ge> n\<close>)
       
  1591       (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps)
       
  1592 qed
       
  1593 
       
  1594 lemma slice1_no_bin [simp]:
       
  1595   "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))"
       
  1596   by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *)
       
  1597 
       
  1598 lemma slice_no_bin [simp]:
       
  1599   "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n)
       
  1600     (bin_to_bl (LENGTH('b::len)) (numeral w)))"
       
  1601   by (simp add: slice_def) (* TODO: neg_numeral *)
       
  1602 
       
  1603 lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))"
       
  1604   by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt)
       
  1605 
       
  1606 lemmas slice_take = slice_take' [unfolded word_size]
       
  1607 
       
  1608 \<comment> \<open>shiftr to a word of the same size is just slice,
       
  1609     slice is just shiftr then ucast\<close>
       
  1610 lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]]
       
  1611 
       
  1612 lemma slice1_down_alt':
       
  1613   "sl = slice1 n w \<Longrightarrow> fs = size sl \<Longrightarrow> fs + k = n \<Longrightarrow>
       
  1614     to_bl sl = takefill False fs (drop k (to_bl w))"
       
  1615   apply (simp add: slice1_eq_of_bl)
       
  1616   apply transfer
       
  1617   apply (simp add: bl_bin_bl_rep_drop)
       
  1618   using drop_takefill
       
  1619   apply force
       
  1620   done
       
  1621 
       
  1622 lemma slice1_up_alt':
       
  1623   "sl = slice1 n w \<Longrightarrow> fs = size sl \<Longrightarrow> fs = n + k \<Longrightarrow>
       
  1624     to_bl sl = takefill False fs (replicate k False @ (to_bl w))"
       
  1625   apply (simp add: slice1_eq_of_bl)
       
  1626   apply transfer
       
  1627   apply (simp add: bl_bin_bl_rep_drop flip: takefill_append)
       
  1628   apply (metis diff_add_inverse)
       
  1629   done
       
  1630 
       
  1631 lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size]
       
  1632 lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size]
       
  1633 lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1]
       
  1634 lemmas slice1_up_alts =
       
  1635   le_add_diff_inverse [symmetric, THEN su1]
       
  1636   le_add_diff_inverse2 [symmetric, THEN su1]
       
  1637 
       
  1638 lemma slice1_tf_tf':
       
  1639   "to_bl (slice1 n w :: 'a::len word) =
       
  1640     rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))"
       
  1641   unfolding slice1_eq_of_bl by (rule word_rev_tf)
       
  1642 
       
  1643 lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric]
       
  1644 
       
  1645 lemma revcast_eq_of_bl:
       
  1646   \<open>(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\<close>
       
  1647   for w :: \<open>'a::len word\<close>
       
  1648   by (simp add: revcast_def slice1_eq_of_bl)
       
  1649 
       
  1650 lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w
       
  1651 
       
  1652 lemma to_bl_revcast:
       
  1653   "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)"
       
  1654   apply (rule nth_equalityI)
       
  1655   apply simp
       
  1656   apply (cases \<open>LENGTH('a) \<le> LENGTH('b)\<close>)
       
  1657    apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff)
       
  1658   done
       
  1659 
       
  1660 lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)"
       
  1661   apply (rule bit_word_eqI)
       
  1662   apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl)
       
  1663   apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl)
       
  1664   done
       
  1665 
       
  1666 lemma of_bl_append:
       
  1667   "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys"
       
  1668   apply transfer
       
  1669   apply (simp add: bl_to_bin_app_cat bin_cat_num)
       
  1670   done
       
  1671 
       
  1672 lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs"
       
  1673   by (rule word_eqI) (auto simp: test_bit_of_bl nth_append)
       
  1674 
       
  1675 lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs"
       
  1676   by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl)
       
  1677 
       
  1678 lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs"
       
  1679   by (cases x) simp_all
       
  1680 
       
  1681 lemma word_split_bl':
       
  1682   "std = size c - size b \<Longrightarrow> (word_split c = (a, b)) \<Longrightarrow>
       
  1683     (a = of_bl (take std (to_bl c)) \<and> b = of_bl (drop std (to_bl c)))"
       
  1684   apply (simp add: word_split_def)
       
  1685   apply transfer
       
  1686   apply (cases \<open>LENGTH('b) \<le> LENGTH('a)\<close>)
       
  1687    apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \<open>LENGTH('a)\<close> \<open>LENGTH('a) - LENGTH('b)\<close> \<open>LENGTH('b)\<close>] min_absorb2)
       
  1688   done
       
  1689 
       
  1690 lemma word_split_bl: "std = size c - size b \<Longrightarrow>
       
  1691     (a = of_bl (take std (to_bl c)) \<and> b = of_bl (drop std (to_bl c))) \<longleftrightarrow>
       
  1692     word_split c = (a, b)"
       
  1693   apply (rule iffI)
       
  1694    defer
       
  1695    apply (erule (1) word_split_bl')
       
  1696   apply (case_tac "word_split c")
       
  1697   apply (auto simp add: word_size)
       
  1698   apply (frule word_split_bl' [rotated])
       
  1699    apply (auto simp add: word_size)
       
  1700   done
       
  1701 
       
  1702 lemma word_split_bl_eq:
       
  1703   "(word_split c :: ('c::len word \<times> 'd::len word)) =
       
  1704     (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)),
       
  1705      of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))"
       
  1706   for c :: "'a::len word"
       
  1707   apply (rule word_split_bl [THEN iffD1])
       
  1708    apply (unfold word_size)
       
  1709    apply (rule refl conjI)+
       
  1710   done
       
  1711 
       
  1712 lemma word_rcat_bl:
       
  1713   \<open>word_rcat wl = of_bl (concat (map to_bl wl))\<close>
       
  1714 proof -
       
  1715   define ws where \<open>ws = rev wl\<close>
       
  1716   moreover have \<open>word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\<close>
       
  1717     apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat)
       
  1718     apply transfer
       
  1719     apply simp
       
  1720     done
       
  1721   ultimately show ?thesis
       
  1722     by simp
       
  1723 qed
       
  1724 
       
  1725 lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)"
       
  1726   by (induct wl) (auto simp: word_size)
       
  1727 
       
  1728 lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size]
       
  1729 
       
  1730 lemma nth_rcat_lem:
       
  1731   "n < length (wl::'a word list) * LENGTH('a::len) \<Longrightarrow>
       
  1732     rev (concat (map to_bl wl)) ! n =
       
  1733     rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))"
       
  1734   apply (induct wl)
       
  1735    apply clarsimp
       
  1736   apply (clarsimp simp add : nth_append size_rcat_lem)
       
  1737   apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less)
       
  1738   apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0)
       
  1739   done
       
  1740 
       
  1741 lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0"
       
  1742   for x :: "'a::comm_monoid_add"
       
  1743   by (induct xs arbitrary: x) (auto simp: add.assoc)
       
  1744 
       
  1745 lemmas word_cat_bl_no_bin [simp] =
       
  1746   word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral]
       
  1747   for a b (* FIXME: negative numerals, 0 and 1 *)
       
  1748 
       
  1749 lemmas word_split_bl_no_bin [simp] =
       
  1750   word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c
       
  1751 
       
  1752 lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq
       
  1753 
       
  1754 lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)"
       
  1755   by (simp add: word_rotl_eq to_bl_use_of_bl)
       
  1756 
       
  1757 lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]]
       
  1758 
       
  1759 lemmas word_rotl_eqs =
       
  1760   blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]]
       
  1761 
       
  1762 lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)"
       
  1763   by (simp add: word_rotr_eq to_bl_use_of_bl)
       
  1764 
       
  1765 lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]]
       
  1766 
       
  1767 lemmas word_rotr_eqs =
       
  1768   brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]]
       
  1769 
       
  1770 declare word_rotr_eqs (1) [simp]
       
  1771 declare word_rotl_eqs (1) [simp]
       
  1772 
       
  1773 lemmas abl_cong = arg_cong [where f = "of_bl"]
       
  1774 
       
  1775 locale word_rotate
       
  1776 begin
       
  1777 
       
  1778 lemmas word_rot_defs' = to_bl_rotl to_bl_rotr
       
  1779 
       
  1780 lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor
       
  1781 
       
  1782 lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]]
       
  1783 
       
  1784 lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2
       
  1785 
       
  1786 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v
       
  1787 
       
  1788 lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map
       
  1789 
       
  1790 end
       
  1791 
       
  1792 lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take,
       
  1793   simplified word_bl_Rep']
       
  1794 
       
  1795 lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take,
       
  1796   simplified word_bl_Rep']
       
  1797 
       
  1798 lemma bl_word_roti_dt':
       
  1799   "n = nat ((- i) mod int (size (w :: 'a::len word))) \<Longrightarrow>
       
  1800     to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)"
       
  1801   apply (unfold word_roti_eq_word_rotr_word_rotl)
       
  1802   apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size)
       
  1803   apply safe
       
  1804    apply (simp add: zmod_zminus1_eq_if)
       
  1805    apply safe
       
  1806     apply (simp add: nat_mult_distrib)
       
  1807    apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj
       
  1808                                       [THEN conjunct2, THEN order_less_imp_le]]
       
  1809                     nat_mod_distrib)
       
  1810   apply (simp add: nat_mod_distrib)
       
  1811   done
       
  1812 
       
  1813 lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size]
       
  1814 
       
  1815 lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]]
       
  1816 lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]]
       
  1817 lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]]
       
  1818 
       
  1819 lemmas word_rotr_dt_no_bin' [simp] =
       
  1820   word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w
       
  1821   (* FIXME: negative numerals, 0 and 1 *)
       
  1822 
       
  1823 lemmas word_rotl_dt_no_bin' [simp] =
       
  1824   word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w
       
  1825   (* FIXME: negative numerals, 0 and 1 *)
       
  1826 
       
  1827 lemma max_word_bl: "to_bl (max_word::'a::len word) = replicate LENGTH('a) True"
       
  1828   by (fact to_bl_n1)
       
  1829 
       
  1830 lemma to_bl_mask:
       
  1831   "to_bl (mask n :: 'a::len word) =
       
  1832   replicate (LENGTH('a) - n) False @
       
  1833     replicate (min (LENGTH('a)) n) True"
       
  1834   by (simp add: mask_bl word_rep_drop min_def)
       
  1835 
       
  1836 lemma map_replicate_True:
       
  1837   "n = length xs \<Longrightarrow>
       
  1838     map (\<lambda>(x,y). x \<and> y) (zip xs (replicate n True)) = xs"
       
  1839   by (induct xs arbitrary: n) auto
       
  1840 
       
  1841 lemma map_replicate_False:
       
  1842   "n = length xs \<Longrightarrow> map (\<lambda>(x,y). x \<and> y)
       
  1843     (zip xs (replicate n False)) = replicate n False"
       
  1844   by (induct xs arbitrary: n) auto
       
  1845 
       
  1846 lemma bl_and_mask:
       
  1847   fixes w :: "'a::len word"
       
  1848     and n :: nat
       
  1849   defines "n' \<equiv> LENGTH('a) - n"
       
  1850   shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)"
       
  1851 proof -
       
  1852   note [simp] = map_replicate_True map_replicate_False
       
  1853   have "to_bl (w AND mask n) = map2 (\<and>) (to_bl w) (to_bl (mask n::'a::len word))"
       
  1854     by (simp add: bl_word_and)
       
  1855   also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)"
       
  1856     by simp
       
  1857   also have "map2 (\<and>) \<dots> (to_bl (mask n::'a::len word)) =
       
  1858       replicate n' False @ drop n' (to_bl w)"
       
  1859     unfolding to_bl_mask n'_def by (subst zip_append) auto
       
  1860   finally show ?thesis .
       
  1861 qed
       
  1862 
       
  1863 lemma drop_rev_takefill:
       
  1864   "length xs \<le> n \<Longrightarrow>
       
  1865     drop (n - length xs) (rev (takefill False n (rev xs))) = xs"
       
  1866   by (simp add: takefill_alt rev_take)
       
  1867 
       
  1868 declare bin_to_bl_def [simp]
       
  1869 
       
  1870 end