| author | wenzelm | 
| Wed, 05 Mar 2008 22:48:50 +0100 | |
| changeset 26205 | 499f08293680 | 
| parent 26086 | 3c243098b64a | 
| child 26557 | 9e7f95903b24 | 
| permissions | -rw-r--r-- | 
| 24333 | 1 | (* | 
| 2 | ID: $Id$ | |
| 3 | Author: Jeremy Dawson, NICTA | |
| 4 | ||
| 5 | contains theorems to do with integers, expressed using Pls, Min, BIT, | |
| 6 | theorems linking them to lists of booleans, and repeated splitting | |
| 7 | and concatenation. | |
| 8 | *) | |
| 9 | ||
| 10 | header "Bool lists and integers" | |
| 11 | ||
| 12 | theory BinBoolList imports BinOperations begin | |
| 13 | ||
| 24465 | 14 | subsection "Arithmetic in terms of bool lists" | 
| 15 | ||
| 16 | consts (* arithmetic operations in terms of the reversed bool list, | |
| 17 | assuming input list(s) the same length, and don't extend them *) | |
| 18 | rbl_succ :: "bool list => bool list" | |
| 19 | rbl_pred :: "bool list => bool list" | |
| 20 | rbl_add :: "bool list => bool list => bool list" | |
| 21 | rbl_mult :: "bool list => bool list => bool list" | |
| 22 | ||
| 23 | primrec | |
| 24 | Nil: "rbl_succ Nil = Nil" | |
| 25 | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" | |
| 26 | ||
| 27 | primrec | |
| 28 | Nil : "rbl_pred Nil = Nil" | |
| 29 | Cons : "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" | |
| 30 | ||
| 31 | primrec (* result is length of first arg, second arg may be longer *) | |
| 32 | Nil : "rbl_add Nil x = Nil" | |
| 33 | Cons : "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in | |
| 34 | (y ~= hd x) # (if hd x & y then rbl_succ ws else ws))" | |
| 35 | ||
| 36 | primrec (* result is length of first arg, second arg may be longer *) | |
| 37 | Nil : "rbl_mult Nil x = Nil" | |
| 38 | Cons : "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in | |
| 39 | if y then rbl_add ws x else ws)" | |
| 24333 | 40 | |
| 41 | lemma tl_take: "tl (take n l) = take (n - 1) (tl l)" | |
| 42 | apply (cases n, clarsimp) | |
| 43 | apply (cases l, auto) | |
| 44 | done | |
| 45 | ||
| 46 | lemma take_butlast [rule_format] : | |
| 47 | "ALL n. n < length l --> take n (butlast l) = take n l" | |
| 48 | apply (induct l, clarsimp) | |
| 49 | apply clarsimp | |
| 50 | apply (case_tac n) | |
| 51 | apply auto | |
| 52 | done | |
| 53 | ||
| 54 | lemma butlast_take [rule_format] : | |
| 55 | "ALL n. n <= length l --> butlast (take n l) = take (n - 1) l" | |
| 56 | apply (induct l, clarsimp) | |
| 57 | apply clarsimp | |
| 58 | apply (case_tac "n") | |
| 59 | apply safe | |
| 60 | apply simp_all | |
| 61 | apply (case_tac "nat") | |
| 62 | apply auto | |
| 63 | done | |
| 64 | ||
| 65 | lemma butlast_drop [rule_format] : | |
| 66 | "ALL n. butlast (drop n l) = drop n (butlast l)" | |
| 67 | apply (induct l) | |
| 68 | apply clarsimp | |
| 69 | apply clarsimp | |
| 70 | apply safe | |
| 71 | apply ((case_tac n, auto)[1])+ | |
| 72 | done | |
| 73 | ||
| 74 | lemma butlast_power: | |
| 75 | "(butlast ^ n) bl = take (length bl - n) bl" | |
| 76 | by (induct n) (auto simp: butlast_take) | |
| 77 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 78 | lemma bin_to_bl_aux_Pls_minus_simp [simp]: | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 79 | "0 < n ==> bin_to_bl_aux n Int.Pls bl = | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 80 | bin_to_bl_aux (n - 1) Int.Pls (False # bl)" | 
| 24333 | 81 | by (cases n) auto | 
| 82 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 83 | lemma bin_to_bl_aux_Min_minus_simp [simp]: | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 84 | "0 < n ==> bin_to_bl_aux n Int.Min bl = | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 85 | bin_to_bl_aux (n - 1) Int.Min (True # bl)" | 
| 24333 | 86 | by (cases n) auto | 
| 87 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 88 | lemma bin_to_bl_aux_Bit_minus_simp [simp]: | 
| 24333 | 89 | "0 < n ==> bin_to_bl_aux n (w BIT b) bl = | 
| 90 | bin_to_bl_aux (n - 1) w ((b = bit.B1) # bl)" | |
| 91 | by (cases n) auto | |
| 92 | ||
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 93 | lemma bin_to_bl_aux_Bit0_minus_simp [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 94 | "0 < n ==> bin_to_bl_aux n (Int.Bit0 w) bl = | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 95 | bin_to_bl_aux (n - 1) w (False # bl)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 96 | by (cases n) auto | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 97 | |
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 98 | lemma bin_to_bl_aux_Bit1_minus_simp [simp]: | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 99 | "0 < n ==> bin_to_bl_aux n (Int.Bit1 w) bl = | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 100 | bin_to_bl_aux (n - 1) w (True # bl)" | 
| 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 101 | by (cases n) auto | 
| 24333 | 102 | |
| 24465 | 103 | (** link between bin and bool list **) | 
| 104 | ||
| 105 | lemma bl_to_bin_aux_append [rule_format] : | |
| 106 | "ALL w. bl_to_bin_aux w (bs @ cs) = bl_to_bin_aux (bl_to_bin_aux w bs) cs" | |
| 107 | by (induct bs) auto | |
| 108 | ||
| 24333 | 109 | lemma bin_to_bl_aux_append [rule_format] : | 
| 110 | "ALL w bs. bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" | |
| 111 | by (induct n) auto | |
| 112 | ||
| 24465 | 113 | lemma bl_to_bin_append: | 
| 114 | "bl_to_bin (bs @ cs) = bl_to_bin_aux (bl_to_bin bs) cs" | |
| 115 | unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) | |
| 116 | ||
| 24333 | 117 | lemma bin_to_bl_aux_alt: | 
| 118 | "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" | |
| 119 | unfolding bin_to_bl_def by (simp add : bin_to_bl_aux_append) | |
| 120 | ||
| 24465 | 121 | lemma bin_to_bl_0: "bin_to_bl 0 bs = []" | 
| 24333 | 122 | unfolding bin_to_bl_def by auto | 
| 123 | ||
| 124 | lemma size_bin_to_bl_aux [rule_format] : | |
| 125 | "ALL w bs. size (bin_to_bl_aux n w bs) = n + length bs" | |
| 126 | by (induct n) auto | |
| 127 | ||
| 24465 | 128 | lemma size_bin_to_bl: "size (bin_to_bl n w) = n" | 
| 24333 | 129 | unfolding bin_to_bl_def by (simp add : size_bin_to_bl_aux) | 
| 130 | ||
| 24465 | 131 | lemma bin_bl_bin' [rule_format] : | 
| 132 | "ALL w bs. bl_to_bin (bin_to_bl_aux n w bs) = | |
| 133 | bl_to_bin_aux (bintrunc n w) bs" | |
| 134 | by (induct n) (auto simp add : bl_to_bin_def) | |
| 135 | ||
| 136 | lemma bin_bl_bin: "bl_to_bin (bin_to_bl n w) = bintrunc n w" | |
| 137 | unfolding bin_to_bl_def bin_bl_bin' by auto | |
| 138 | ||
| 139 | lemma bl_bin_bl' [rule_format] : | |
| 140 | "ALL w n. bin_to_bl (n + length bs) (bl_to_bin_aux w bs) = | |
| 141 | bin_to_bl_aux n w bs" | |
| 142 | apply (induct "bs") | |
| 143 | apply auto | |
| 144 | apply (simp_all only : add_Suc [symmetric]) | |
| 145 | apply (auto simp add : bin_to_bl_def) | |
| 146 | done | |
| 147 | ||
| 148 | lemma bl_bin_bl: "bin_to_bl (length bs) (bl_to_bin bs) = bs" | |
| 149 | unfolding bl_to_bin_def | |
| 150 | apply (rule box_equals) | |
| 151 | apply (rule bl_bin_bl') | |
| 152 | prefer 2 | |
| 153 | apply (rule bin_to_bl_aux.Z) | |
| 154 | apply simp | |
| 155 | done | |
| 156 | ||
| 157 | declare | |
| 158 | bin_to_bl_0 [simp] | |
| 159 | size_bin_to_bl [simp] | |
| 160 | bin_bl_bin [simp] | |
| 161 | bl_bin_bl [simp] | |
| 162 | ||
| 163 | lemma bl_to_bin_inj: | |
| 164 | "bl_to_bin bs = bl_to_bin cs ==> length bs = length cs ==> bs = cs" | |
| 165 | apply (rule_tac box_equals) | |
| 166 | defer | |
| 167 | apply (rule bl_bin_bl) | |
| 168 | apply (rule bl_bin_bl) | |
| 169 | apply simp | |
| 170 | done | |
| 171 | ||
| 172 | lemma bl_to_bin_False: "bl_to_bin (False # bl) = bl_to_bin bl" | |
| 173 | unfolding bl_to_bin_def by auto | |
| 174 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 175 | lemma bl_to_bin_Nil: "bl_to_bin [] = Int.Pls" | 
| 24465 | 176 | unfolding bl_to_bin_def by auto | 
| 177 | ||
| 24333 | 178 | lemma bin_to_bl_Pls_aux [rule_format] : | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 179 | "ALL bl. bin_to_bl_aux n Int.Pls bl = replicate n False @ bl" | 
| 24333 | 180 | by (induct n) (auto simp: replicate_app_Cons_same) | 
| 181 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 182 | lemma bin_to_bl_Pls: "bin_to_bl n Int.Pls = replicate n False" | 
| 24333 | 183 | unfolding bin_to_bl_def by (simp add : bin_to_bl_Pls_aux) | 
| 184 | ||
| 185 | lemma bin_to_bl_Min_aux [rule_format] : | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 186 | "ALL bl. bin_to_bl_aux n Int.Min bl = replicate n True @ bl" | 
| 24333 | 187 | by (induct n) (auto simp: replicate_app_Cons_same) | 
| 188 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 189 | lemma bin_to_bl_Min: "bin_to_bl n Int.Min = replicate n True" | 
| 24333 | 190 | unfolding bin_to_bl_def by (simp add : bin_to_bl_Min_aux) | 
| 191 | ||
| 24465 | 192 | lemma bl_to_bin_rep_F: | 
| 193 | "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" | |
| 194 | apply (simp add: bin_to_bl_Pls_aux [symmetric] bin_bl_bin') | |
| 195 | apply (simp add: bl_to_bin_def) | |
| 196 | done | |
| 197 | ||
| 198 | lemma bin_to_bl_trunc: | |
| 199 | "n <= m ==> bin_to_bl n (bintrunc m w) = bin_to_bl n w" | |
| 200 | by (auto intro: bl_to_bin_inj) | |
| 201 | ||
| 202 | declare | |
| 203 | bin_to_bl_trunc [simp] | |
| 204 | bl_to_bin_False [simp] | |
| 205 | bl_to_bin_Nil [simp] | |
| 206 | ||
| 24333 | 207 | lemma bin_to_bl_aux_bintr [rule_format] : | 
| 208 | "ALL m bin bl. bin_to_bl_aux n (bintrunc m bin) bl = | |
| 209 | replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" | |
| 210 | apply (induct_tac "n") | |
| 211 | apply clarsimp | |
| 212 | apply clarsimp | |
| 213 | apply (case_tac "m") | |
| 214 | apply (clarsimp simp: bin_to_bl_Pls_aux) | |
| 215 | apply (erule thin_rl) | |
| 216 | apply (induct_tac n) | |
| 217 | apply auto | |
| 218 | done | |
| 219 | ||
| 220 | lemmas bin_to_bl_bintr = | |
| 221 | bin_to_bl_aux_bintr [where bl = "[]", folded bin_to_bl_def] | |
| 222 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 223 | lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = Int.Pls" | 
| 24465 | 224 | by (induct n) auto | 
| 225 | ||
| 24333 | 226 | lemma len_bin_to_bl_aux [rule_format] : | 
| 227 | "ALL w bs. length (bin_to_bl_aux n w bs) = n + length bs" | |
| 228 | by (induct "n") auto | |
| 229 | ||
| 230 | lemma len_bin_to_bl [simp]: "length (bin_to_bl n w) = n" | |
| 231 | unfolding bin_to_bl_def len_bin_to_bl_aux by auto | |
| 232 | ||
| 233 | lemma sign_bl_bin' [rule_format] : | |
| 234 | "ALL w. bin_sign (bl_to_bin_aux w bs) = bin_sign w" | |
| 235 | by (induct bs) auto | |
| 236 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 237 | lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = Int.Pls" | 
| 24333 | 238 | unfolding bl_to_bin_def by (simp add : sign_bl_bin') | 
| 239 | ||
| 240 | lemma bl_sbin_sign_aux [rule_format] : | |
| 241 | "ALL w bs. hd (bin_to_bl_aux (Suc n) w bs) = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 242 | (bin_sign (sbintrunc n w) = Int.Min)" | 
| 24333 | 243 | apply (induct "n") | 
| 244 | apply clarsimp | |
| 245 | apply (case_tac w rule: bin_exhaust) | |
| 246 | apply (simp split add : bit.split) | |
| 247 | apply clarsimp | |
| 248 | done | |
| 249 | ||
| 250 | lemma bl_sbin_sign: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 251 | "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = Int.Min)" | 
| 24333 | 252 | unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) | 
| 253 | ||
| 254 | lemma bin_nth_of_bl_aux [rule_format] : | |
| 255 | "ALL w. bin_nth (bl_to_bin_aux w bl) n = | |
| 256 | (n < size bl & rev bl ! n | n >= length bl & bin_nth w (n - size bl))" | |
| 257 | apply (induct_tac bl) | |
| 258 | apply clarsimp | |
| 259 | apply clarsimp | |
| 260 | apply (cut_tac x=n and y="size list" in linorder_less_linear) | |
| 261 | apply (erule disjE, simp add: nth_append)+ | |
| 262 | apply (simp add: nth_append) | |
| 263 | done | |
| 264 | ||
| 265 | lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl & rev bl ! n)"; | |
| 266 | unfolding bl_to_bin_def by (simp add : bin_nth_of_bl_aux) | |
| 267 | ||
| 268 | lemma bin_nth_bl [rule_format] : "ALL m w. n < m --> | |
| 269 | bin_nth w n = nth (rev (bin_to_bl m w)) n" | |
| 270 | apply (induct n) | |
| 271 | apply clarsimp | |
| 272 | apply (case_tac m, clarsimp) | |
| 273 | apply (clarsimp simp: bin_to_bl_def) | |
| 274 | apply (simp add: bin_to_bl_aux_alt) | |
| 275 | apply clarsimp | |
| 276 | apply (case_tac m, clarsimp) | |
| 277 | apply (clarsimp simp: bin_to_bl_def) | |
| 278 | apply (simp add: bin_to_bl_aux_alt) | |
| 279 | done | |
| 280 | ||
| 24465 | 281 | lemma nth_rev [rule_format] : | 
| 282 | "n < length xs --> rev xs ! n = xs ! (length xs - 1 - n)" | |
| 283 | apply (induct_tac "xs") | |
| 284 | apply simp | |
| 285 | apply (clarsimp simp add : nth_append nth.simps split add : nat.split) | |
| 286 | apply (rule_tac f = "%n. list ! n" in arg_cong) | |
| 287 | apply arith | |
| 288 | done | |
| 289 | ||
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 290 | lemmas nth_rev_alt = nth_rev [where xs = "rev ys", simplified, standard] | 
| 24465 | 291 | |
| 24333 | 292 | lemma nth_bin_to_bl_aux [rule_format] : | 
| 293 | "ALL w n bl. n < m + length bl --> (bin_to_bl_aux m w bl) ! n = | |
| 294 | (if n < m then bin_nth w (m - 1 - n) else bl ! (n - m))" | |
| 295 | apply (induct_tac "m") | |
| 296 | apply clarsimp | |
| 297 | apply clarsimp | |
| 298 | apply (case_tac w rule: bin_exhaust) | |
| 299 | apply clarsimp | |
| 300 | apply (case_tac "na - n") | |
| 301 | apply arith | |
| 302 | apply simp | |
| 303 | apply (rule_tac f = "%n. bl ! n" in arg_cong) | |
| 304 | apply arith | |
| 305 | done | |
| 306 | ||
| 307 | lemma nth_bin_to_bl: "n < m ==> (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" | |
| 308 | unfolding bin_to_bl_def by (simp add : nth_bin_to_bl_aux) | |
| 309 | ||
| 310 | lemma bl_to_bin_lt2p_aux [rule_format] : | |
| 311 | "ALL w. bl_to_bin_aux w bs < (w + 1) * (2 ^ length bs)" | |
| 312 | apply (induct "bs") | |
| 313 | apply clarsimp | |
| 314 | apply clarsimp | |
| 315 | apply safe | |
| 316 | apply (erule allE, erule xtr8 [rotated], | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 317 | simp add: numeral_simps ring_simps cong add : number_of_False_cong)+ | 
| 24333 | 318 | done | 
| 319 | ||
| 320 | lemma bl_to_bin_lt2p: "bl_to_bin bs < (2 ^ length bs)" | |
| 321 | apply (unfold bl_to_bin_def) | |
| 322 | apply (rule xtr1) | |
| 323 | prefer 2 | |
| 324 | apply (rule bl_to_bin_lt2p_aux) | |
| 325 | apply simp | |
| 326 | done | |
| 327 | ||
| 328 | lemma bl_to_bin_ge2p_aux [rule_format] : | |
| 329 | "ALL w. bl_to_bin_aux w bs >= w * (2 ^ length bs)" | |
| 330 | apply (induct bs) | |
| 331 | apply clarsimp | |
| 332 | apply clarsimp | |
| 333 | apply safe | |
| 334 | apply (erule allE, erule less_eq_less.order_trans [rotated], | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 335 | simp add: numeral_simps ring_simps cong add : number_of_False_cong)+ | 
| 24333 | 336 | done | 
| 337 | ||
| 338 | lemma bl_to_bin_ge0: "bl_to_bin bs >= 0" | |
| 339 | apply (unfold bl_to_bin_def) | |
| 340 | apply (rule xtr4) | |
| 341 | apply (rule bl_to_bin_ge2p_aux) | |
| 342 | apply simp | |
| 343 | done | |
| 344 | ||
| 345 | lemma butlast_rest_bin: | |
| 346 | "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" | |
| 347 | apply (unfold bin_to_bl_def) | |
| 348 | apply (cases w rule: bin_exhaust) | |
| 349 | apply (cases n, clarsimp) | |
| 350 | apply clarsimp | |
| 351 | apply (auto simp add: bin_to_bl_aux_alt) | |
| 352 | done | |
| 353 | ||
| 354 | lemmas butlast_bin_rest = butlast_rest_bin | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 355 | [where w="bl_to_bin bl" and n="length bl", simplified, standard] | 
| 24333 | 356 | |
| 357 | lemma butlast_rest_bl2bin_aux [rule_format] : | |
| 358 | "ALL w. bl ~= [] --> | |
| 359 | bl_to_bin_aux w (butlast bl) = bin_rest (bl_to_bin_aux w bl)" | |
| 360 | by (induct bl) auto | |
| 361 | ||
| 362 | lemma butlast_rest_bl2bin: | |
| 363 | "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" | |
| 364 | apply (unfold bl_to_bin_def) | |
| 365 | apply (cases bl) | |
| 366 | apply (auto simp add: butlast_rest_bl2bin_aux) | |
| 367 | done | |
| 368 | ||
| 369 | lemma trunc_bl2bin_aux [rule_format] : | |
| 370 | "ALL w. bintrunc m (bl_to_bin_aux w bl) = | |
| 371 | bl_to_bin_aux (bintrunc (m - length bl) w) (drop (length bl - m) bl)" | |
| 372 | apply (induct_tac bl) | |
| 373 | apply clarsimp | |
| 374 | apply clarsimp | |
| 375 | apply safe | |
| 376 | apply (case_tac "m - size list") | |
| 377 | apply (simp add : diff_is_0_eq [THEN iffD1, THEN Suc_diff_le]) | |
| 378 | apply simp | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 379 | apply (rule_tac f = "%nat. bl_to_bin_aux (Int.Bit1 (bintrunc nat w)) list" | 
| 24333 | 380 | in arg_cong) | 
| 381 | apply simp | |
| 382 | apply (case_tac "m - size list") | |
| 383 | apply (simp add: diff_is_0_eq [THEN iffD1, THEN Suc_diff_le]) | |
| 384 | apply simp | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26008diff
changeset | 385 | apply (rule_tac f = "%nat. bl_to_bin_aux (Int.Bit0 (bintrunc nat w)) list" | 
| 24333 | 386 | in arg_cong) | 
| 387 | apply simp | |
| 388 | done | |
| 389 | ||
| 390 | lemma trunc_bl2bin: | |
| 391 | "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" | |
| 392 | unfolding bl_to_bin_def by (simp add : trunc_bl2bin_aux) | |
| 393 | ||
| 394 | lemmas trunc_bl2bin_len [simp] = | |
| 395 | trunc_bl2bin [of "length bl" bl, simplified, standard] | |
| 396 | ||
| 397 | lemma bl2bin_drop: | |
| 398 | "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" | |
| 399 | apply (rule trans) | |
| 400 | prefer 2 | |
| 401 | apply (rule trunc_bl2bin [symmetric]) | |
| 402 | apply (cases "k <= length bl") | |
| 403 | apply auto | |
| 404 | done | |
| 405 | ||
| 406 | lemma nth_rest_power_bin [rule_format] : | |
| 407 | "ALL n. bin_nth ((bin_rest ^ k) w) n = bin_nth w (n + k)" | |
| 408 | apply (induct k, clarsimp) | |
| 409 | apply clarsimp | |
| 410 | apply (simp only: bin_nth.Suc [symmetric] add_Suc) | |
| 411 | done | |
| 412 | ||
| 413 | lemma take_rest_power_bin: | |
| 414 | "m <= n ==> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^ (n - m)) w)" | |
| 415 | apply (rule nth_equalityI) | |
| 416 | apply simp | |
| 417 | apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) | |
| 418 | done | |
| 419 | ||
| 24465 | 420 | lemma hd_butlast: "size xs > 1 ==> hd (butlast xs) = hd xs" | 
| 421 | by (cases xs) auto | |
| 24333 | 422 | |
| 423 | lemma last_bin_last' [rule_format] : | |
| 424 | "ALL w. size xs > 0 --> last xs = (bin_last (bl_to_bin_aux w xs) = bit.B1)" | |
| 425 | by (induct xs) auto | |
| 426 | ||
| 427 | lemma last_bin_last: | |
| 428 | "size xs > 0 ==> last xs = (bin_last (bl_to_bin xs) = bit.B1)" | |
| 429 | unfolding bl_to_bin_def by (erule last_bin_last') | |
| 430 | ||
| 431 | lemma bin_last_last: | |
| 432 | "bin_last w = (if last (bin_to_bl (Suc n) w) then bit.B1 else bit.B0)" | |
| 433 | apply (unfold bin_to_bl_def) | |
| 434 | apply simp | |
| 435 | apply (auto simp add: bin_to_bl_aux_alt) | |
| 436 | done | |
| 437 | ||
| 24465 | 438 | (** links between bit-wise operations and operations on bool lists **) | 
| 439 | ||
| 24333 | 440 | lemma app2_Nil [simp]: "app2 f [] ys = []" | 
| 441 | unfolding app2_def by auto | |
| 442 | ||
| 443 | lemma app2_Cons [simp]: | |
| 444 | "app2 f (x # xs) (y # ys) = f x y # app2 f xs ys" | |
| 445 | unfolding app2_def by auto | |
| 446 | ||
| 447 | lemma bl_xor_aux_bin [rule_format] : "ALL v w bs cs. | |
| 448 | app2 (%x y. x ~= y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | |
| 24353 | 449 | bin_to_bl_aux n (v XOR w) (app2 (%x y. x ~= y) bs cs)" | 
| 24333 | 450 | apply (induct_tac n) | 
| 451 | apply safe | |
| 452 | apply simp | |
| 453 | apply (case_tac v rule: bin_exhaust) | |
| 454 | apply (case_tac w rule: bin_exhaust) | |
| 455 | apply clarsimp | |
| 456 | apply (case_tac b) | |
| 457 | apply (case_tac ba, safe, simp_all)+ | |
| 458 | done | |
| 459 | ||
| 460 | lemma bl_or_aux_bin [rule_format] : "ALL v w bs cs. | |
| 461 | app2 (op | ) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | |
| 24353 | 462 | bin_to_bl_aux n (v OR w) (app2 (op | ) bs cs)" | 
| 24333 | 463 | apply (induct_tac n) | 
| 464 | apply safe | |
| 465 | apply simp | |
| 466 | apply (case_tac v rule: bin_exhaust) | |
| 467 | apply (case_tac w rule: bin_exhaust) | |
| 468 | apply clarsimp | |
| 469 | apply (case_tac b) | |
| 470 | apply (case_tac ba, safe, simp_all)+ | |
| 471 | done | |
| 472 | ||
| 473 | lemma bl_and_aux_bin [rule_format] : "ALL v w bs cs. | |
| 474 | app2 (op & ) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = | |
| 24353 | 475 | bin_to_bl_aux n (v AND w) (app2 (op & ) bs cs)" | 
| 24333 | 476 | apply (induct_tac n) | 
| 477 | apply safe | |
| 478 | apply simp | |
| 479 | apply (case_tac v rule: bin_exhaust) | |
| 480 | apply (case_tac w rule: bin_exhaust) | |
| 481 | apply clarsimp | |
| 482 | apply (case_tac b) | |
| 483 | apply (case_tac ba, safe, simp_all)+ | |
| 484 | done | |
| 485 | ||
| 486 | lemma bl_not_aux_bin [rule_format] : | |
| 487 | "ALL w cs. map Not (bin_to_bl_aux n w cs) = | |
| 24353 | 488 | bin_to_bl_aux n (NOT w) (map Not cs)" | 
| 24333 | 489 | apply (induct n) | 
| 490 | apply clarsimp | |
| 491 | apply clarsimp | |
| 492 | apply (case_tac w rule: bin_exhaust) | |
| 493 | apply (case_tac b) | |
| 494 | apply auto | |
| 495 | done | |
| 496 | ||
| 497 | lemmas bl_not_bin = bl_not_aux_bin | |
| 498 | [where cs = "[]", unfolded bin_to_bl_def [symmetric] map.simps] | |
| 499 | ||
| 500 | lemmas bl_and_bin = bl_and_aux_bin [where bs="[]" and cs="[]", | |
| 501 | unfolded app2_Nil, folded bin_to_bl_def] | |
| 502 | ||
| 503 | lemmas bl_or_bin = bl_or_aux_bin [where bs="[]" and cs="[]", | |
| 504 | unfolded app2_Nil, folded bin_to_bl_def] | |
| 505 | ||
| 506 | lemmas bl_xor_bin = bl_xor_aux_bin [where bs="[]" and cs="[]", | |
| 507 | unfolded app2_Nil, folded bin_to_bl_def] | |
| 508 | ||
| 509 | lemma drop_bin2bl_aux [rule_format] : | |
| 510 | "ALL m bin bs. drop m (bin_to_bl_aux n bin bs) = | |
| 511 | bin_to_bl_aux (n - m) bin (drop (m - n) bs)" | |
| 512 | apply (induct n, clarsimp) | |
| 513 | apply clarsimp | |
| 514 | apply (case_tac bin rule: bin_exhaust) | |
| 515 | apply (case_tac "m <= n", simp) | |
| 516 | apply (case_tac "m - n", simp) | |
| 517 | apply simp | |
| 518 | apply (rule_tac f = "%nat. drop nat bs" in arg_cong) | |
| 519 | apply simp | |
| 520 | done | |
| 521 | ||
| 522 | lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" | |
| 523 | unfolding bin_to_bl_def by (simp add : drop_bin2bl_aux) | |
| 524 | ||
| 525 | lemma take_bin2bl_lem1 [rule_format] : | |
| 526 | "ALL w bs. take m (bin_to_bl_aux m w bs) = bin_to_bl m w" | |
| 527 | apply (induct m, clarsimp) | |
| 528 | apply clarsimp | |
| 529 | apply (simp add: bin_to_bl_aux_alt) | |
| 530 | apply (simp add: bin_to_bl_def) | |
| 531 | apply (simp add: bin_to_bl_aux_alt) | |
| 532 | done | |
| 533 | ||
| 534 | lemma take_bin2bl_lem [rule_format] : | |
| 535 | "ALL w bs. take m (bin_to_bl_aux (m + n) w bs) = | |
| 536 | take m (bin_to_bl (m + n) w)" | |
| 537 | apply (induct n) | |
| 538 | apply clarify | |
| 539 | apply (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1) | |
| 540 | apply simp | |
| 541 | done | |
| 542 | ||
| 543 | lemma bin_split_take [rule_format] : | |
| 544 | "ALL b c. bin_split n c = (a, b) --> | |
| 545 | bin_to_bl m a = take m (bin_to_bl (m + n) c)" | |
| 546 | apply (induct n) | |
| 547 | apply clarsimp | |
| 548 | apply (clarsimp simp: Let_def split: ls_splits) | |
| 549 | apply (simp add: bin_to_bl_def) | |
| 550 | apply (simp add: take_bin2bl_lem) | |
| 551 | done | |
| 552 | ||
| 553 | lemma bin_split_take1: | |
| 554 | "k = m + n ==> bin_split n c = (a, b) ==> | |
| 555 | bin_to_bl m a = take m (bin_to_bl k c)" | |
| 556 | by (auto elim: bin_split_take) | |
| 557 | ||
| 558 | lemma nth_takefill [rule_format] : "ALL m l. m < n --> | |
| 559 | takefill fill n l ! m = (if m < length l then l ! m else fill)" | |
| 560 | apply (induct n, clarsimp) | |
| 561 | apply clarsimp | |
| 562 | apply (case_tac m) | |
| 563 | apply (simp split: list.split) | |
| 564 | apply clarsimp | |
| 565 | apply (erule allE)+ | |
| 566 | apply (erule (1) impE) | |
| 567 | apply (simp split: list.split) | |
| 568 | done | |
| 569 | ||
| 570 | lemma takefill_alt [rule_format] : | |
| 571 | "ALL l. takefill fill n l = take n l @ replicate (n - length l) fill" | |
| 572 | by (induct n) (auto split: list.split) | |
| 573 | ||
| 574 | lemma takefill_replicate [simp]: | |
| 575 | "takefill fill n (replicate m fill) = replicate n fill" | |
| 576 | by (simp add : takefill_alt replicate_add [symmetric]) | |
| 577 | ||
| 578 | lemma takefill_le' [rule_format] : | |
| 579 | "ALL l n. n = m + k --> takefill x m (takefill x n l) = takefill x m l" | |
| 580 | by (induct m) (auto split: list.split) | |
| 581 | ||
| 582 | lemma length_takefill [simp]: "length (takefill fill n l) = n" | |
| 583 | by (simp add : takefill_alt) | |
| 584 | ||
| 585 | lemma take_takefill': | |
| 586 | "!!w n. n = k + m ==> take k (takefill fill n w) = takefill fill k w" | |
| 587 | by (induct k) (auto split add : list.split) | |
| 588 | ||
| 589 | lemma drop_takefill: | |
| 590 | "!!w. drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" | |
| 591 | by (induct k) (auto split add : list.split) | |
| 592 | ||
| 593 | lemma takefill_le [simp]: | |
| 594 | "m \<le> n \<Longrightarrow> takefill x m (takefill x n l) = takefill x m l" | |
| 595 | by (auto simp: le_iff_add takefill_le') | |
| 596 | ||
| 597 | lemma take_takefill [simp]: | |
| 598 | "m \<le> n \<Longrightarrow> take m (takefill fill n w) = takefill fill m w" | |
| 599 | by (auto simp: le_iff_add take_takefill') | |
| 600 | ||
| 601 | lemma takefill_append: | |
| 602 | "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" | |
| 603 | by (induct xs) auto | |
| 604 | ||
| 605 | lemma takefill_same': | |
| 606 | "l = length xs ==> takefill fill l xs = xs" | |
| 607 | by clarify (induct xs, auto) | |
| 608 | ||
| 609 | lemmas takefill_same [simp] = takefill_same' [OF refl] | |
| 610 | ||
| 611 | lemma takefill_bintrunc: | |
| 612 | "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" | |
| 613 | apply (rule nth_equalityI) | |
| 614 | apply simp | |
| 615 | apply (clarsimp simp: nth_takefill nth_rev nth_bin_to_bl bin_nth_of_bl) | |
| 616 | done | |
| 617 | ||
| 618 | lemma bl_bin_bl_rtf: | |
| 619 | "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" | |
| 620 | by (simp add : takefill_bintrunc) | |
| 621 | ||
| 622 | lemmas bl_bin_bl_rep_drop = | |
| 623 | bl_bin_bl_rtf [simplified takefill_alt, | |
| 624 | simplified, simplified rev_take, simplified] | |
| 625 | ||
| 626 | lemma tf_rev: | |
| 627 | "n + k = m + length bl ==> takefill x m (rev (takefill y n bl)) = | |
| 628 | rev (takefill y m (rev (takefill x k (rev bl))))" | |
| 629 | apply (rule nth_equalityI) | |
| 630 | apply (auto simp add: nth_takefill nth_rev) | |
| 631 | apply (rule_tac f = "%n. bl ! n" in arg_cong) | |
| 632 | apply arith | |
| 633 | done | |
| 634 | ||
| 635 | lemma takefill_minus: | |
| 636 | "0 < n ==> takefill fill (Suc (n - 1)) w = takefill fill n w" | |
| 637 | by auto | |
| 638 | ||
| 639 | lemmas takefill_Suc_cases = | |
| 640 | list.cases [THEN takefill.Suc [THEN trans], standard] | |
| 641 | ||
| 642 | lemmas takefill_Suc_Nil = takefill_Suc_cases (1) | |
| 643 | lemmas takefill_Suc_Cons = takefill_Suc_cases (2) | |
| 644 | ||
| 645 | lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] | |
| 646 | takefill_minus [symmetric, THEN trans], standard] | |
| 647 | ||
| 648 | lemmas takefill_pred_simps [simp] = | |
| 649 | takefill_minus_simps [where n="number_of bin", simplified nobm1, standard] | |
| 650 | ||
| 651 | (* links with function bl_to_bin *) | |
| 652 | ||
| 653 | lemma bl_to_bin_aux_cat: | |
| 654 | "!!nv v. bl_to_bin_aux (bin_cat w nv v) bs = | |
| 655 | bin_cat w (nv + length bs) (bl_to_bin_aux v bs)" | |
| 656 | apply (induct bs) | |
| 657 | apply simp | |
| 658 | apply (simp add: bin_cat_Suc_Bit [symmetric] del: bin_cat.simps) | |
| 659 | done | |
| 660 | ||
| 661 | lemma bin_to_bl_aux_cat: | |
| 662 | "!!w bs. bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = | |
| 663 | bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" | |
| 664 | by (induct nw) auto | |
| 665 | ||
| 666 | lemmas bl_to_bin_aux_alt = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 667 | bl_to_bin_aux_cat [where nv = "0" and v = "Int.Pls", | 
| 24333 | 668 | simplified bl_to_bin_def [symmetric], simplified] | 
| 669 | ||
| 670 | lemmas bin_to_bl_cat = | |
| 671 | bin_to_bl_aux_cat [where bs = "[]", folded bin_to_bl_def] | |
| 672 | ||
| 673 | lemmas bl_to_bin_aux_app_cat = | |
| 674 | trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] | |
| 675 | ||
| 676 | lemmas bin_to_bl_aux_cat_app = | |
| 677 | trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] | |
| 678 | ||
| 679 | lemmas bl_to_bin_app_cat = bl_to_bin_aux_app_cat | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 680 | [where w = "Int.Pls", folded bl_to_bin_def] | 
| 24333 | 681 | |
| 682 | lemmas bin_to_bl_cat_app = bin_to_bl_aux_cat_app | |
| 683 | [where bs = "[]", folded bin_to_bl_def] | |
| 684 | ||
| 685 | (* bl_to_bin_app_cat_alt and bl_to_bin_app_cat are easily interderivable *) | |
| 686 | lemma bl_to_bin_app_cat_alt: | |
| 687 | "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" | |
| 688 | by (simp add : bl_to_bin_app_cat) | |
| 689 | ||
| 690 | lemma mask_lem: "(bl_to_bin (True # replicate n False)) = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 691 | Int.succ (bl_to_bin (replicate n True))" | 
| 24333 | 692 | apply (unfold bl_to_bin_def) | 
| 693 | apply (induct n) | |
| 694 | apply simp | |
| 695 | apply (simp only: Suc_eq_add_numeral_1 replicate_add | |
| 696 | append_Cons [symmetric] bl_to_bin_aux_append) | |
| 697 | apply simp | |
| 698 | done | |
| 699 | ||
| 24465 | 700 | (* function bl_of_nth *) | 
| 24333 | 701 | lemma length_bl_of_nth [simp]: "length (bl_of_nth n f) = n" | 
| 702 | by (induct n) auto | |
| 703 | ||
| 704 | lemma nth_bl_of_nth [simp]: | |
| 705 | "m < n \<Longrightarrow> rev (bl_of_nth n f) ! m = f m" | |
| 706 | apply (induct n) | |
| 707 | apply simp | |
| 708 | apply (clarsimp simp add : nth_append) | |
| 709 | apply (rule_tac f = "f" in arg_cong) | |
| 710 | apply simp | |
| 711 | done | |
| 712 | ||
| 713 | lemma bl_of_nth_inj: | |
| 714 | "(!!k. k < n ==> f k = g k) ==> bl_of_nth n f = bl_of_nth n g" | |
| 715 | by (induct n) auto | |
| 716 | ||
| 717 | lemma bl_of_nth_nth_le [rule_format] : "ALL xs. | |
| 718 | length xs >= n --> bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs"; | |
| 719 | apply (induct n, clarsimp) | |
| 720 | apply clarsimp | |
| 721 | apply (rule trans [OF _ hd_Cons_tl]) | |
| 722 | apply (frule Suc_le_lessD) | |
| 723 | apply (simp add: nth_rev trans [OF drop_Suc drop_tl, symmetric]) | |
| 724 | apply (subst hd_drop_conv_nth) | |
| 725 | apply force | |
| 726 | apply simp_all | |
| 727 | apply (rule_tac f = "%n. drop n xs" in arg_cong) | |
| 728 | apply simp | |
| 729 | done | |
| 730 | ||
| 731 | lemmas bl_of_nth_nth [simp] = order_refl [THEN bl_of_nth_nth_le, simplified] | |
| 732 | ||
| 733 | lemma size_rbl_pred: "length (rbl_pred bl) = length bl" | |
| 734 | by (induct bl) auto | |
| 735 | ||
| 736 | lemma size_rbl_succ: "length (rbl_succ bl) = length bl" | |
| 737 | by (induct bl) auto | |
| 738 | ||
| 739 | lemma size_rbl_add: | |
| 740 | "!!cl. length (rbl_add bl cl) = length bl" | |
| 741 | by (induct bl) (auto simp: Let_def size_rbl_succ) | |
| 742 | ||
| 743 | lemma size_rbl_mult: | |
| 744 | "!!cl. length (rbl_mult bl cl) = length bl" | |
| 745 | by (induct bl) (auto simp add : Let_def size_rbl_add) | |
| 746 | ||
| 747 | lemmas rbl_sizes [simp] = | |
| 748 | size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult | |
| 749 | ||
| 750 | lemmas rbl_Nils = | |
| 751 | rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil | |
| 752 | ||
| 753 | lemma rbl_pred: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 754 | "!!bin. rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (Int.pred bin))" | 
| 24333 | 755 | apply (induct n, simp) | 
| 756 | apply (unfold bin_to_bl_def) | |
| 757 | apply clarsimp | |
| 758 | apply (case_tac bin rule: bin_exhaust) | |
| 759 | apply (case_tac b) | |
| 760 | apply (clarsimp simp: bin_to_bl_aux_alt)+ | |
| 761 | done | |
| 762 | ||
| 763 | lemma rbl_succ: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 764 | "!!bin. rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (Int.succ bin))" | 
| 24333 | 765 | apply (induct n, simp) | 
| 766 | apply (unfold bin_to_bl_def) | |
| 767 | apply clarsimp | |
| 768 | apply (case_tac bin rule: bin_exhaust) | |
| 769 | apply (case_tac b) | |
| 770 | apply (clarsimp simp: bin_to_bl_aux_alt)+ | |
| 771 | done | |
| 772 | ||
| 773 | lemma rbl_add: | |
| 774 | "!!bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = | |
| 775 | rev (bin_to_bl n (bina + binb))" | |
| 776 | apply (induct n, simp) | |
| 777 | apply (unfold bin_to_bl_def) | |
| 778 | apply clarsimp | |
| 779 | apply (case_tac bina rule: bin_exhaust) | |
| 780 | apply (case_tac binb rule: bin_exhaust) | |
| 781 | apply (case_tac b) | |
| 782 | apply (case_tac [!] "ba") | |
| 783 | apply (auto simp: rbl_succ succ_def bin_to_bl_aux_alt Let_def add_ac) | |
| 784 | done | |
| 785 | ||
| 786 | lemma rbl_add_app2: | |
| 787 | "!!blb. length blb >= length bla ==> | |
| 788 | rbl_add bla (blb @ blc) = rbl_add bla blb" | |
| 789 | apply (induct bla, simp) | |
| 790 | apply clarsimp | |
| 791 | apply (case_tac blb, clarsimp) | |
| 792 | apply (clarsimp simp: Let_def) | |
| 793 | done | |
| 794 | ||
| 795 | lemma rbl_add_take2: | |
| 796 | "!!blb. length blb >= length bla ==> | |
| 797 | rbl_add bla (take (length bla) blb) = rbl_add bla blb" | |
| 798 | apply (induct bla, simp) | |
| 799 | apply clarsimp | |
| 800 | apply (case_tac blb, clarsimp) | |
| 801 | apply (clarsimp simp: Let_def) | |
| 802 | done | |
| 803 | ||
| 804 | lemma rbl_add_long: | |
| 805 | "m >= n ==> rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = | |
| 806 | rev (bin_to_bl n (bina + binb))" | |
| 807 | apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) | |
| 808 | apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) | |
| 809 | apply (rule rev_swap [THEN iffD1]) | |
| 810 | apply (simp add: rev_take drop_bin2bl) | |
| 811 | apply simp | |
| 812 | done | |
| 813 | ||
| 814 | lemma rbl_mult_app2: | |
| 815 | "!!blb. length blb >= length bla ==> | |
| 816 | rbl_mult bla (blb @ blc) = rbl_mult bla blb" | |
| 817 | apply (induct bla, simp) | |
| 818 | apply clarsimp | |
| 819 | apply (case_tac blb, clarsimp) | |
| 820 | apply (clarsimp simp: Let_def rbl_add_app2) | |
| 821 | done | |
| 822 | ||
| 823 | lemma rbl_mult_take2: | |
| 824 | "length blb >= length bla ==> | |
| 825 | rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" | |
| 826 | apply (rule trans) | |
| 827 | apply (rule rbl_mult_app2 [symmetric]) | |
| 828 | apply simp | |
| 829 | apply (rule_tac f = "rbl_mult bla" in arg_cong) | |
| 830 | apply (rule append_take_drop_id) | |
| 831 | done | |
| 832 | ||
| 833 | lemma rbl_mult_gt1: | |
| 834 | "m >= length bl ==> rbl_mult bl (rev (bin_to_bl m binb)) = | |
| 835 | rbl_mult bl (rev (bin_to_bl (length bl) binb))" | |
| 836 | apply (rule trans) | |
| 837 | apply (rule rbl_mult_take2 [symmetric]) | |
| 838 | apply simp_all | |
| 839 | apply (rule_tac f = "rbl_mult bl" in arg_cong) | |
| 840 | apply (rule rev_swap [THEN iffD1]) | |
| 841 | apply (simp add: rev_take drop_bin2bl) | |
| 842 | done | |
| 843 | ||
| 844 | lemma rbl_mult_gt: | |
| 845 | "m > n ==> rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = | |
| 846 | rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" | |
| 847 | by (auto intro: trans [OF rbl_mult_gt1]) | |
| 848 | ||
| 849 | lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] | |
| 850 | ||
| 851 | lemma rbbl_Cons: | |
| 852 | "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (x BIT If b bit.B1 bit.B0))" | |
| 853 | apply (unfold bin_to_bl_def) | |
| 854 | apply simp | |
| 855 | apply (simp add: bin_to_bl_aux_alt) | |
| 856 | done | |
| 857 | ||
| 858 | lemma rbl_mult: "!!bina binb. | |
| 859 | rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = | |
| 860 | rev (bin_to_bl n (bina * binb))" | |
| 861 | apply (induct n) | |
| 862 | apply simp | |
| 863 | apply (unfold bin_to_bl_def) | |
| 864 | apply clarsimp | |
| 865 | apply (case_tac bina rule: bin_exhaust) | |
| 866 | apply (case_tac binb rule: bin_exhaust) | |
| 867 | apply (case_tac b) | |
| 868 | apply (case_tac [!] "ba") | |
| 869 | apply (auto simp: bin_to_bl_aux_alt Let_def) | |
| 870 | apply (auto simp: rbbl_Cons rbl_mult_Suc rbl_add) | |
| 871 | done | |
| 872 | ||
| 873 | lemma rbl_add_split: | |
| 874 | "P (rbl_add (y # ys) (x # xs)) = | |
| 875 | (ALL ws. length ws = length ys --> ws = rbl_add ys xs --> | |
| 26008 | 876 | (y --> ((x --> P (False # rbl_succ ws)) & (~ x --> P (True # ws)))) & | 
| 24333 | 877 | (~ y --> P (x # ws)))" | 
| 878 | apply (auto simp add: Let_def) | |
| 879 | apply (case_tac [!] "y") | |
| 880 | apply auto | |
| 881 | done | |
| 882 | ||
| 883 | lemma rbl_mult_split: | |
| 884 | "P (rbl_mult (y # ys) xs) = | |
| 885 | (ALL ws. length ws = Suc (length ys) --> ws = False # rbl_mult ys xs --> | |
| 886 | (y --> P (rbl_add ws xs)) & (~ y --> P ws))" | |
| 887 | by (clarsimp simp add : Let_def) | |
| 888 | ||
| 889 | lemma and_len: "xs = ys ==> xs = ys & length xs = length ys" | |
| 890 | by auto | |
| 891 | ||
| 892 | lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" | |
| 893 | by auto | |
| 894 | ||
| 895 | lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" | |
| 896 | by auto | |
| 897 | ||
| 898 | lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" | |
| 899 | by auto | |
| 900 | ||
| 24465 | 901 | lemma if_Not_x: "(if p then ~ x else x) = (p = (~ x))" | 
| 902 | by auto | |
| 903 | ||
| 904 | lemma if_x_Not: "(if p then x else ~ x) = (p = x)" | |
| 905 | by auto | |
| 906 | ||
| 24333 | 907 | lemma if_same_and: "(If p x y & If p u v) = (if p then x & u else y & v)" | 
| 908 | by auto | |
| 909 | ||
| 910 | lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = (u) else y = (v))" | |
| 911 | by auto | |
| 912 | ||
| 913 | lemma if_same_eq_not: | |
| 914 | "(If p x y = (~ If p u v)) = (if p then x = (~u) else y = (~v))" | |
| 915 | by auto | |
| 916 | ||
| 917 | (* note - if_Cons can cause blowup in the size, if p is complex, | |
| 918 | so make a simproc *) | |
| 919 | lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" | |
| 920 | by auto | |
| 921 | ||
| 922 | lemma if_single: | |
| 923 | "(if xc then [xab] else [an]) = [if xc then xab else an]" | |
| 924 | by auto | |
| 925 | ||
| 24465 | 926 | lemma if_bool_simps: | 
| 927 | "If p True y = (p | y) & If p False y = (~p & y) & | |
| 928 | If p y True = (p --> y) & If p y False = (p & y)" | |
| 929 | by auto | |
| 930 | ||
| 931 | lemmas if_simps = if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps | |
| 932 | ||
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 933 | lemmas seqr = eq_reflection [where x = "size w", standard] | 
| 24333 | 934 | |
| 935 | lemmas tl_Nil = tl.simps (1) | |
| 936 | lemmas tl_Cons = tl.simps (2) | |
| 937 | ||
| 938 | ||
| 24350 | 939 | subsection "Repeated splitting or concatenation" | 
| 24333 | 940 | |
| 941 | lemma sclem: | |
| 942 | "size (concat (map (bin_to_bl n) xs)) = length xs * n" | |
| 943 | by (induct xs) auto | |
| 944 | ||
| 945 | lemma bin_cat_foldl_lem [rule_format] : | |
| 946 | "ALL x. foldl (%u. bin_cat u n) x xs = | |
| 947 | bin_cat x (size xs * n) (foldl (%u. bin_cat u n) y xs)" | |
| 948 | apply (induct xs) | |
| 949 | apply simp | |
| 950 | apply clarify | |
| 951 | apply (simp (no_asm)) | |
| 952 | apply (frule asm_rl) | |
| 953 | apply (drule spec) | |
| 954 | apply (erule trans) | |
| 955 | apply (drule_tac x = "bin_cat y n a" in spec) | |
| 956 | apply (simp add : bin_cat_assoc_sym min_def) | |
| 957 | done | |
| 958 | ||
| 959 | lemma bin_rcat_bl: | |
| 960 | "(bin_rcat n wl) = bl_to_bin (concat (map (bin_to_bl n) wl))" | |
| 961 | apply (unfold bin_rcat_def) | |
| 962 | apply (rule sym) | |
| 963 | apply (induct wl) | |
| 964 | apply (auto simp add : bl_to_bin_append) | |
| 965 | apply (simp add : bl_to_bin_aux_alt sclem) | |
| 966 | apply (simp add : bin_cat_foldl_lem [symmetric]) | |
| 967 | done | |
| 968 | ||
| 969 | lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps | |
| 970 | lemmas rsplit_aux_simps = bin_rsplit_aux_simps | |
| 971 | ||
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 972 | lemmas th_if_simp1 = split_if [where P = "op = l", | 
| 24333 | 973 | THEN iffD1, THEN conjunct1, THEN mp, standard] | 
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 974 | lemmas th_if_simp2 = split_if [where P = "op = l", | 
| 24333 | 975 | THEN iffD1, THEN conjunct2, THEN mp, standard] | 
| 976 | ||
| 977 | lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] | |
| 978 | ||
| 979 | lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] | |
| 980 | (* these safe to [simp add] as require calculating m - n *) | |
| 981 | lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] | |
| 982 | lemmas rbscl = bin_rsplit_aux_simp2s (2) | |
| 983 | ||
| 984 | lemmas rsplit_aux_0_simps [simp] = | |
| 985 | rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] | |
| 986 | ||
| 987 | lemma bin_rsplit_aux_append: | |
| 988 | "bin_rsplit_aux (n, bs @ cs, m, c) = bin_rsplit_aux (n, bs, m, c) @ cs" | |
| 989 | apply (rule_tac u=n and v=bs and w=m and x=c in bin_rsplit_aux.induct) | |
| 990 | apply (subst bin_rsplit_aux.simps) | |
| 991 | apply (subst bin_rsplit_aux.simps) | |
| 992 | apply (clarsimp split: ls_splits) | |
| 993 | done | |
| 994 | ||
| 995 | lemma bin_rsplitl_aux_append: | |
| 996 | "bin_rsplitl_aux (n, bs @ cs, m, c) = bin_rsplitl_aux (n, bs, m, c) @ cs" | |
| 997 | apply (rule_tac u=n and v=bs and w=m and x=c in bin_rsplitl_aux.induct) | |
| 998 | apply (subst bin_rsplitl_aux.simps) | |
| 999 | apply (subst bin_rsplitl_aux.simps) | |
| 1000 | apply (clarsimp split: ls_splits) | |
| 1001 | done | |
| 1002 | ||
| 1003 | lemmas rsplit_aux_apps [where bs = "[]"] = | |
| 1004 | bin_rsplit_aux_append bin_rsplitl_aux_append | |
| 1005 | ||
| 1006 | lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def | |
| 1007 | ||
| 1008 | lemmas rsplit_aux_alts = rsplit_aux_apps | |
| 1009 | [unfolded append_Nil rsplit_def_auxs [symmetric]] | |
| 1010 | ||
| 1011 | lemma bin_split_minus: "0 < n ==> bin_split (Suc (n - 1)) w = bin_split n w" | |
| 1012 | by auto | |
| 1013 | ||
| 1014 | lemmas bin_split_minus_simp = | |
| 1015 | bin_split.Suc [THEN [2] bin_split_minus [symmetric, THEN trans], standard] | |
| 1016 | ||
| 1017 | lemma bin_split_pred_simp [simp]: | |
| 1018 | "(0::nat) < number_of bin \<Longrightarrow> | |
| 1019 | bin_split (number_of bin) w = | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25350diff
changeset | 1020 | (let (w1, w2) = bin_split (number_of (Int.pred bin)) (bin_rest w) | 
| 24333 | 1021 | in (w1, w2 BIT bin_last w))" | 
| 1022 | by (simp only: nobm1 bin_split_minus_simp) | |
| 1023 | ||
| 24465 | 1024 | declare bin_split_pred_simp [simp] | 
| 1025 | ||
| 24333 | 1026 | lemma bin_rsplit_aux_simp_alt: | 
| 1027 | "bin_rsplit_aux (n, bs, m, c) = | |
| 1028 | (if m = 0 \<or> n = 0 | |
| 1029 | then bs | |
| 1030 | else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" | |
| 1031 | apply (rule trans) | |
| 1032 | apply (subst bin_rsplit_aux.simps, rule refl) | |
| 1033 | apply (simp add: rsplit_aux_alts) | |
| 1034 | done | |
| 1035 | ||
| 1036 | lemmas bin_rsplit_simp_alt = | |
| 1037 | trans [OF bin_rsplit_def [THEN meta_eq_to_obj_eq] | |
| 1038 | bin_rsplit_aux_simp_alt, standard] | |
| 1039 | ||
| 1040 | lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] | |
| 1041 | ||
| 1042 | lemma bin_rsplit_size_sign' [rule_format] : | |
| 1043 | "n > 0 ==> (ALL nw w. rev sw = bin_rsplit n (nw, w) --> | |
| 1044 | (ALL v: set sw. bintrunc n v = v))" | |
| 1045 | apply (induct sw) | |
| 1046 | apply clarsimp | |
| 1047 | apply clarsimp | |
| 1048 | apply (drule bthrs) | |
| 1049 | apply (simp (no_asm_use) add: Let_def split: ls_splits) | |
| 1050 | apply clarify | |
| 1051 | apply (erule impE, rule exI, erule exI) | |
| 1052 | apply (drule split_bintrunc) | |
| 1053 | apply simp | |
| 1054 | done | |
| 1055 | ||
| 1056 | lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl | |
| 1057 | rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]], | |
| 1058 | standard] | |
| 1059 | ||
| 1060 | lemma bin_nth_rsplit [rule_format] : | |
| 1061 | "n > 0 ==> m < n ==> (ALL w k nw. rev sw = bin_rsplit n (nw, w) --> | |
| 1062 | k < size sw --> bin_nth (sw ! k) m = bin_nth w (k * n + m))" | |
| 1063 | apply (induct sw) | |
| 1064 | apply clarsimp | |
| 1065 | apply clarsimp | |
| 1066 | apply (drule bthrs) | |
| 1067 | apply (simp (no_asm_use) add: Let_def split: ls_splits) | |
| 1068 | apply clarify | |
| 1069 | apply (erule allE, erule impE, erule exI) | |
| 1070 | apply (case_tac k) | |
| 1071 | apply clarsimp | |
| 1072 | prefer 2 | |
| 1073 | apply clarsimp | |
| 1074 | apply (erule allE) | |
| 1075 | apply (erule (1) impE) | |
| 1076 | apply (drule bin_nth_split, erule conjE, erule allE, | |
| 1077 | erule trans, simp add : add_ac)+ | |
| 1078 | done | |
| 1079 | ||
| 1080 | lemma bin_rsplit_all: | |
| 1081 | "0 < nw ==> nw <= n ==> bin_rsplit n (nw, w) = [bintrunc n w]" | |
| 1082 | unfolding bin_rsplit_def | |
| 1083 | by (clarsimp dest!: split_bintrunc simp: rsplit_aux_simp2ls split: ls_splits) | |
| 1084 | ||
| 1085 | lemma bin_rsplit_l [rule_format] : | |
| 1086 | "ALL bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" | |
| 1087 | apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) | |
| 1088 | apply (simp (no_asm) add : bin_rsplitl_def bin_rsplit_def) | |
| 1089 | apply (rule allI) | |
| 1090 | apply (subst bin_rsplitl_aux.simps) | |
| 1091 | apply (subst bin_rsplit_aux.simps) | |
| 1092 | apply (clarsimp simp: rsplit_aux_alts Let_def split: ls_splits) | |
| 1093 | apply (drule bin_split_trunc) | |
| 1094 | apply (drule sym [THEN trans], assumption) | |
| 1095 | apply fast | |
| 1096 | done | |
| 1097 | ||
| 1098 | lemma bin_rsplit_rcat [rule_format] : | |
| 1099 | "n > 0 --> bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" | |
| 1100 | apply (unfold bin_rsplit_def bin_rcat_def) | |
| 1101 | apply (rule_tac xs = "ws" in rev_induct) | |
| 1102 | apply clarsimp | |
| 1103 | apply clarsimp | |
| 1104 | apply (clarsimp simp add: bin_split_cat rsplit_aux_alts) | |
| 1105 | done | |
| 1106 | ||
| 1107 | lemma bin_rsplit_aux_len_le [rule_format] : | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1108 | "ALL ws m. n \<noteq> 0 --> ws = bin_rsplit_aux (n, bs, nw, w) --> | 
| 24333 | 1109 | (length ws <= m) = (nw + length bs * n <= m * n)" | 
| 1110 | apply (rule_tac u=n and v=bs and w=nw and x=w in bin_rsplit_aux.induct) | |
| 1111 | apply (subst bin_rsplit_aux.simps) | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1112 | apply (simp add:lrlem Let_def split: ls_splits ) | 
| 24333 | 1113 | done | 
| 1114 | ||
| 1115 | lemma bin_rsplit_len_le: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1116 | "n \<noteq> 0 --> ws = bin_rsplit n (nw, w) --> (length ws <= m) = (nw <= m * n)" | 
| 24333 | 1117 | unfolding bin_rsplit_def by (clarsimp simp add : bin_rsplit_aux_len_le) | 
| 1118 | ||
| 1119 | lemma bin_rsplit_aux_len [rule_format] : | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1120 | "n\<noteq>0 --> length (bin_rsplit_aux (n, cs, nw, w)) = | 
| 24333 | 1121 | (nw + n - 1) div n + length cs" | 
| 1122 | apply (rule_tac u=n and v=cs and w=nw and x=w in bin_rsplit_aux.induct) | |
| 1123 | apply (subst bin_rsplit_aux.simps) | |
| 1124 | apply (clarsimp simp: Let_def split: ls_splits) | |
| 1125 | apply (erule thin_rl) | |
| 1126 | apply (case_tac "m <= n") | |
| 1127 | prefer 2 | |
| 1128 | apply (simp add: div_add_self2 [symmetric]) | |
| 1129 | apply (case_tac m, clarsimp) | |
| 1130 | apply (simp add: div_add_self2) | |
| 1131 | done | |
| 1132 | ||
| 1133 | lemma bin_rsplit_len: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1134 | "n\<noteq>0 ==> length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" | 
| 24333 | 1135 | unfolding bin_rsplit_def by (clarsimp simp add : bin_rsplit_aux_len) | 
| 1136 | ||
| 1137 | lemma bin_rsplit_aux_len_indep [rule_format] : | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1138 | "n\<noteq>0 ==> (ALL v bs. length bs = length cs --> | 
| 24333 | 1139 | length (bin_rsplit_aux (n, bs, nw, v)) = | 
| 1140 | length (bin_rsplit_aux (n, cs, nw, w)))" | |
| 1141 | apply (rule_tac u=n and v=cs and w=nw and x=w in bin_rsplit_aux.induct) | |
| 1142 | apply clarsimp | |
| 1143 | apply (simp (no_asm_simp) add: bin_rsplit_aux_simp_alt Let_def | |
| 1144 | split: ls_splits) | |
| 1145 | apply clarify | |
| 1146 | apply (erule allE)+ | |
| 1147 | apply (erule impE) | |
| 1148 | apply (fast elim!: sym) | |
| 1149 | apply (simp (no_asm_use) add: rsplit_aux_alts) | |
| 1150 | apply (erule impE) | |
| 1151 | apply (rule_tac x="ba # bs" in exI) | |
| 1152 | apply auto | |
| 1153 | done | |
| 1154 | ||
| 1155 | lemma bin_rsplit_len_indep: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1156 | "n\<noteq>0 ==> length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" | 
| 24333 | 1157 | apply (unfold bin_rsplit_def) | 
| 1158 | apply (erule bin_rsplit_aux_len_indep) | |
| 1159 | apply (rule refl) | |
| 1160 | done | |
| 1161 | ||
| 1162 | end | |
| 1163 |