| author | wenzelm | 
| Thu, 21 Aug 2008 20:53:31 +0200 | |
| changeset 27943 | f34ff5e7728f | 
| parent 27682 | 25aceefd4786 | 
| child 28059 | 295a8fc92684 | 
| permissions | -rw-r--r-- | 
| 24333 | 1 | (* | 
| 2 | ID: $Id$ | |
| 3 | Author: Jeremy Dawson and Gerwin Klein, NICTA | |
| 4 | ||
| 5 | contains arithmetic theorems for word, instantiations to | |
| 6 | arithmetic type classes and tactics for reducing word arithmetic | |
| 7 | to linear arithmetic on int or nat | |
| 8 | *) | |
| 9 | ||
| 24350 | 10 | header {* Word Arithmetic *}
 | 
| 11 | ||
| 26560 | 12 | theory WordArith | 
| 13 | imports WordDefinition | |
| 14 | begin | |
| 24465 | 15 | |
| 16 | lemma word_less_alt: "(a < b) = (uint a < uint b)" | |
| 17 | unfolding word_less_def word_le_def | |
| 18 | by (auto simp del: word_uint.Rep_inject | |
| 19 | simp: word_uint.Rep_inject [symmetric]) | |
| 20 | ||
| 21 | lemma signed_linorder: "linorder word_sle word_sless" | |
| 22 | apply unfold_locales | |
| 23 | apply (unfold word_sle_def word_sless_def) | |
| 24 | by auto | |
| 25 | ||
| 26 | interpretation signed: linorder ["word_sle" "word_sless"] | |
| 27 | by (rule signed_linorder) | |
| 28 | ||
| 25762 | 29 | lemmas word_arith_wis = | 
| 24333 | 30 | word_add_def word_mult_def word_minus_def | 
| 31 | word_succ_def word_pred_def word_0_wi word_1_wi | |
| 32 | ||
| 24465 | 33 | lemma udvdI: | 
| 34 | "0 \<le> n ==> uint b = n * uint a ==> a udvd b" | |
| 35 | by (auto simp: udvd_def) | |
| 36 | ||
| 37 | lemmas word_div_no [simp] = | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 38 | word_div_def [of "number_of a" "number_of b", standard] | 
| 24465 | 39 | |
| 40 | lemmas word_mod_no [simp] = | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 41 | word_mod_def [of "number_of a" "number_of b", standard] | 
| 24465 | 42 | |
| 43 | lemmas word_less_no [simp] = | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 44 | word_less_def [of "number_of a" "number_of b", standard] | 
| 24465 | 45 | |
| 46 | lemmas word_le_no [simp] = | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 47 | word_le_def [of "number_of a" "number_of b", standard] | 
| 24465 | 48 | |
| 49 | lemmas word_sless_no [simp] = | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 50 | word_sless_def [of "number_of a" "number_of b", standard] | 
| 24465 | 51 | |
| 52 | lemmas word_sle_no [simp] = | |
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 53 | word_sle_def [of "number_of a" "number_of b", standard] | 
| 24465 | 54 | |
| 24333 | 55 | (* following two are available in class number_ring, | 
| 56 | but convenient to have them here here; | |
| 57 | note - the number_ring versions, numeral_0_eq_0 and numeral_1_eq_1 | |
| 58 | are in the default simpset, so to use the automatic simplifications for | |
| 59 | (eg) sint (number_of bin) on sint 1, must do | |
| 60 | (simp add: word_1_no del: numeral_1_eq_1) | |
| 61 | *) | |
| 62 | lemmas word_0_wi_Pls = word_0_wi [folded Pls_def] | |
| 63 | lemmas word_0_no = word_0_wi_Pls [folded word_no_wi] | |
| 64 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 65 | lemma int_one_bin: "(1 :: int) == (Int.Pls BIT bit.B1)" | 
| 24333 | 66 | unfolding Pls_def Bit_def by auto | 
| 67 | ||
| 68 | lemma word_1_no: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 69 | "(1 :: 'a :: len0 word) == number_of (Int.Pls BIT bit.B1)" | 
| 24333 | 70 | unfolding word_1_wi word_number_of_def int_one_bin by auto | 
| 71 | ||
| 72 | lemma word_m1_wi: "-1 == word_of_int -1" | |
| 73 | by (rule word_number_of_alt) | |
| 74 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 75 | lemma word_m1_wi_Min: "-1 = word_of_int Int.Min" | 
| 24333 | 76 | by (simp add: word_m1_wi number_of_eq) | 
| 77 | ||
| 24465 | 78 | lemma word_0_bl: "of_bl [] = 0" | 
| 79 | unfolding word_0_wi of_bl_def by (simp add : Pls_def) | |
| 80 | ||
| 81 | lemma word_1_bl: "of_bl [True] = 1" | |
| 82 | unfolding word_1_wi of_bl_def | |
| 83 | by (simp add : bl_to_bin_def Bit_def Pls_def) | |
| 84 | ||
| 24333 | 85 | lemma uint_0 [simp] : "(uint 0 = 0)" | 
| 86 | unfolding word_0_wi | |
| 87 | by (simp add: word_ubin.eq_norm Pls_def [symmetric]) | |
| 88 | ||
| 24465 | 89 | lemma of_bl_0 [simp] : "of_bl (replicate n False) = 0" | 
| 90 | by (simp add : word_0_wi of_bl_def bl_to_bin_rep_False Pls_def) | |
| 91 | ||
| 92 | lemma to_bl_0: | |
| 93 |   "to_bl (0::'a::len0 word) = replicate (len_of TYPE('a)) False"
 | |
| 94 | unfolding uint_bl | |
| 95 | by (simp add : word_size bin_to_bl_Pls Pls_def [symmetric]) | |
| 96 | ||
| 24333 | 97 | lemma uint_0_iff: "(uint x = 0) = (x = 0)" | 
| 98 | by (auto intro!: word_uint.Rep_eqD) | |
| 99 | ||
| 100 | lemma unat_0_iff: "(unat x = 0) = (x = 0)" | |
| 101 | unfolding unat_def by (auto simp add : nat_eq_iff uint_0_iff) | |
| 102 | ||
| 103 | lemma unat_0 [simp]: "unat 0 = 0" | |
| 104 | unfolding unat_def by auto | |
| 105 | ||
| 24465 | 106 | lemma size_0_same': "size w = 0 ==> w = (v :: 'a :: len0 word)" | 
| 24333 | 107 | apply (unfold word_size) | 
| 108 | apply (rule box_equals) | |
| 109 | defer | |
| 110 | apply (rule word_uint.Rep_inverse)+ | |
| 111 | apply (rule word_ubin.norm_eq_iff [THEN iffD1]) | |
| 112 | apply simp | |
| 113 | done | |
| 114 | ||
| 115 | lemmas size_0_same = size_0_same' [folded word_size] | |
| 116 | ||
| 117 | lemmas unat_eq_0 = unat_0_iff | |
| 118 | lemmas unat_eq_zero = unat_0_iff | |
| 119 | ||
| 120 | lemma unat_gt_0: "(0 < unat x) = (x ~= 0)" | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 121 | by (auto simp: unat_0_iff [symmetric]) | 
| 24333 | 122 | |
| 123 | lemma ucast_0 [simp] : "ucast 0 = 0" | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 124 | unfolding ucast_def | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 125 | by simp (simp add: word_0_wi) | 
| 24333 | 126 | |
| 127 | lemma sint_0 [simp] : "sint 0 = 0" | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 128 | unfolding sint_uint | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 129 | by (simp add: Pls_def [symmetric]) | 
| 24333 | 130 | |
| 131 | lemma scast_0 [simp] : "scast 0 = 0" | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 132 | apply (unfold scast_def) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 133 | apply simp | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 134 | apply (simp add: word_0_wi) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 135 | done | 
| 24333 | 136 | |
| 137 | lemma sint_n1 [simp] : "sint -1 = -1" | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 138 | apply (unfold word_m1_wi_Min) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 139 | apply (simp add: word_sbin.eq_norm) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 140 | apply (unfold Min_def number_of_eq) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 141 | apply simp | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 142 | done | 
| 24333 | 143 | |
| 144 | lemma scast_n1 [simp] : "scast -1 = -1" | |
| 145 | apply (unfold scast_def sint_n1) | |
| 146 | apply (unfold word_number_of_alt) | |
| 147 | apply (rule refl) | |
| 148 | done | |
| 149 | ||
| 24465 | 150 | lemma uint_1 [simp] : "uint (1 :: 'a :: len word) = 1" | 
| 24333 | 151 | unfolding word_1_wi | 
| 152 | by (simp add: word_ubin.eq_norm int_one_bin bintrunc_minus_simps) | |
| 153 | ||
| 24465 | 154 | lemma unat_1 [simp] : "unat (1 :: 'a :: len word) = 1" | 
| 24333 | 155 | by (unfold unat_def uint_1) auto | 
| 156 | ||
| 24465 | 157 | lemma ucast_1 [simp] : "ucast (1 :: 'a :: len word) = 1" | 
| 24333 | 158 | unfolding ucast_def word_1_wi | 
| 159 | by (simp add: word_ubin.eq_norm int_one_bin bintrunc_minus_simps) | |
| 160 | ||
| 161 | (* abstraction preserves the operations | |
| 162 | (the definitions tell this for bins in range uint) *) | |
| 163 | ||
| 164 | lemmas arths = | |
| 165 | bintr_ariths [THEN word_ubin.norm_eq_iff [THEN iffD1], | |
| 166 | folded word_ubin.eq_norm, standard] | |
| 167 | ||
| 168 | lemma wi_homs: | |
| 169 | shows | |
| 170 | wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and | |
| 171 | wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and | |
| 172 | wi_hom_neg: "- word_of_int a = word_of_int (- a)" and | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 173 | wi_hom_succ: "word_succ (word_of_int a) = word_of_int (Int.succ a)" and | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 174 | wi_hom_pred: "word_pred (word_of_int a) = word_of_int (Int.pred a)" | 
| 24333 | 175 | by (auto simp: word_arith_wis arths) | 
| 176 | ||
| 177 | lemmas wi_hom_syms = wi_homs [symmetric] | |
| 178 | ||
| 24465 | 179 | lemma word_sub_def: "a - b == a + - (b :: 'a :: len0 word)" | 
| 180 | unfolding word_sub_wi diff_def | |
| 181 | by (simp only : word_uint.Rep_inverse wi_hom_syms) | |
| 24333 | 182 | |
| 183 | lemmas word_diff_minus = word_sub_def [THEN meta_eq_to_obj_eq, standard] | |
| 184 | ||
| 185 | lemma word_of_int_sub_hom: | |
| 186 | "(word_of_int a) - word_of_int b = word_of_int (a - b)" | |
| 187 | unfolding word_sub_def diff_def by (simp only : wi_homs) | |
| 188 | ||
| 189 | lemmas new_word_of_int_homs = | |
| 190 | word_of_int_sub_hom wi_homs word_0_wi word_1_wi | |
| 191 | ||
| 192 | lemmas new_word_of_int_hom_syms = new_word_of_int_homs [symmetric, standard] | |
| 193 | ||
| 194 | lemmas word_of_int_hom_syms = | |
| 195 | new_word_of_int_hom_syms [unfolded succ_def pred_def] | |
| 196 | ||
| 197 | lemmas word_of_int_homs = | |
| 198 | new_word_of_int_homs [unfolded succ_def pred_def] | |
| 199 | ||
| 200 | lemmas word_of_int_add_hom = word_of_int_homs (2) | |
| 201 | lemmas word_of_int_mult_hom = word_of_int_homs (3) | |
| 202 | lemmas word_of_int_minus_hom = word_of_int_homs (4) | |
| 203 | lemmas word_of_int_succ_hom = word_of_int_homs (5) | |
| 204 | lemmas word_of_int_pred_hom = word_of_int_homs (6) | |
| 205 | lemmas word_of_int_0_hom = word_of_int_homs (7) | |
| 206 | lemmas word_of_int_1_hom = word_of_int_homs (8) | |
| 207 | ||
| 208 | (* now, to get the weaker results analogous to word_div/mod_def *) | |
| 209 | ||
| 210 | lemmas word_arith_alts = | |
| 25762 | 211 | word_sub_wi [unfolded succ_def pred_def, standard] | 
| 24333 | 212 | word_arith_wis [unfolded succ_def pred_def, standard] | 
| 213 | ||
| 214 | lemmas word_sub_alt = word_arith_alts (1) | |
| 215 | lemmas word_add_alt = word_arith_alts (2) | |
| 216 | lemmas word_mult_alt = word_arith_alts (3) | |
| 217 | lemmas word_minus_alt = word_arith_alts (4) | |
| 218 | lemmas word_succ_alt = word_arith_alts (5) | |
| 219 | lemmas word_pred_alt = word_arith_alts (6) | |
| 220 | lemmas word_0_alt = word_arith_alts (7) | |
| 221 | lemmas word_1_alt = word_arith_alts (8) | |
| 222 | ||
| 24350 | 223 | subsection "Transferring goals from words to ints" | 
| 24333 | 224 | |
| 225 | lemma word_ths: | |
| 226 | shows | |
| 227 | word_succ_p1: "word_succ a = a + 1" and | |
| 228 | word_pred_m1: "word_pred a = a - 1" and | |
| 229 | word_pred_succ: "word_pred (word_succ a) = a" and | |
| 230 | word_succ_pred: "word_succ (word_pred a) = a" and | |
| 231 | word_mult_succ: "word_succ a * b = b + a * b" | |
| 232 | by (rule word_uint.Abs_cases [of b], | |
| 233 | rule word_uint.Abs_cases [of a], | |
| 234 | simp add: pred_def succ_def add_commute mult_commute | |
| 235 | ring_distribs new_word_of_int_homs)+ | |
| 236 | ||
| 237 | lemmas uint_cong = arg_cong [where f = uint] | |
| 238 | ||
| 239 | lemmas uint_word_ariths = | |
| 240 | word_arith_alts [THEN trans [OF uint_cong int_word_uint], standard] | |
| 241 | ||
| 242 | lemmas uint_word_arith_bintrs = uint_word_ariths [folded bintrunc_mod2p] | |
| 243 | ||
| 244 | (* similar expressions for sint (arith operations) *) | |
| 245 | lemmas sint_word_ariths = uint_word_arith_bintrs | |
| 246 | [THEN uint_sint [symmetric, THEN trans], | |
| 247 | unfolded uint_sint bintr_arith1s bintr_ariths | |
| 24465 | 248 | len_gt_0 [THEN bin_sbin_eq_iff'] word_sbin.norm_Rep, standard] | 
| 249 | ||
| 250 | lemmas uint_div_alt = word_div_def | |
| 25762 | 251 | [THEN trans [OF uint_cong int_word_uint], standard] | 
| 24465 | 252 | lemmas uint_mod_alt = word_mod_def | 
| 25762 | 253 | [THEN trans [OF uint_cong int_word_uint], standard] | 
| 24333 | 254 | |
| 255 | lemma word_pred_0_n1: "word_pred 0 = word_of_int -1" | |
| 256 | unfolding word_pred_def number_of_eq | |
| 257 | by (simp add : pred_def word_no_wi) | |
| 258 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 259 | lemma word_pred_0_Min: "word_pred 0 = word_of_int Int.Min" | 
| 24333 | 260 | by (simp add: word_pred_0_n1 number_of_eq) | 
| 261 | ||
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 262 | lemma word_m1_Min: "- 1 = word_of_int Int.Min" | 
| 24333 | 263 | unfolding Min_def by (simp only: word_of_int_hom_syms) | 
| 264 | ||
| 265 | lemma succ_pred_no [simp]: | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 266 | "word_succ (number_of bin) = number_of (Int.succ bin) & | 
| 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 267 | word_pred (number_of bin) = number_of (Int.pred bin)" | 
| 24333 | 268 | unfolding word_number_of_def by (simp add : new_word_of_int_homs) | 
| 269 | ||
| 270 | lemma word_sp_01 [simp] : | |
| 271 | "word_succ -1 = 0 & word_succ 0 = 1 & word_pred 0 = -1 & word_pred 1 = 0" | |
| 272 | by (unfold word_0_no word_1_no) auto | |
| 273 | ||
| 274 | (* alternative approach to lifting arithmetic equalities *) | |
| 275 | lemma word_of_int_Ex: | |
| 276 | "\<exists>y. x = word_of_int y" | |
| 277 | by (rule_tac x="uint x" in exI) simp | |
| 278 | ||
| 24465 | 279 | lemma word_arith_eqs: | 
| 280 | fixes a :: "'a::len0 word" | |
| 281 | fixes b :: "'a::len0 word" | |
| 282 | shows | |
| 283 | word_add_0: "0 + a = a" and | |
| 284 | word_add_0_right: "a + 0 = a" and | |
| 285 | word_mult_1: "1 * a = a" and | |
| 286 | word_mult_1_right: "a * 1 = a" and | |
| 287 | word_add_commute: "a + b = b + a" and | |
| 288 | word_add_assoc: "a + b + c = a + (b + c)" and | |
| 289 | word_add_left_commute: "a + (b + c) = b + (a + c)" and | |
| 290 | word_mult_commute: "a * b = b * a" and | |
| 291 | word_mult_assoc: "a * b * c = a * (b * c)" and | |
| 292 | word_mult_left_commute: "a * (b * c) = b * (a * c)" and | |
| 293 | word_left_distrib: "(a + b) * c = a * c + b * c" and | |
| 294 | word_right_distrib: "a * (b + c) = a * b + a * c" and | |
| 295 | word_left_minus: "- a + a = 0" and | |
| 296 | word_diff_0_right: "a - 0 = a" and | |
| 297 | word_diff_self: "a - a = 0" | |
| 298 | using word_of_int_Ex [of a] | |
| 299 | word_of_int_Ex [of b] | |
| 300 | word_of_int_Ex [of c] | |
| 301 | by (auto simp: word_of_int_hom_syms [symmetric] | |
| 302 | zadd_0_right add_commute add_assoc add_left_commute | |
| 303 | mult_commute mult_assoc mult_left_commute | |
| 304 | plus_times.left_distrib plus_times.right_distrib) | |
| 305 | ||
| 306 | lemmas word_add_ac = word_add_commute word_add_assoc word_add_left_commute | |
| 307 | lemmas word_mult_ac = word_mult_commute word_mult_assoc word_mult_left_commute | |
| 308 | ||
| 309 | lemmas word_plus_ac0 = word_add_0 word_add_0_right word_add_ac | |
| 310 | lemmas word_times_ac1 = word_mult_1 word_mult_1_right word_mult_ac | |
| 311 | ||
| 312 | ||
| 24350 | 313 | subsection "Order on fixed-length words" | 
| 24333 | 314 | |
| 24465 | 315 | lemma word_order_trans: "x <= y ==> y <= z ==> x <= (z :: 'a :: len0 word)" | 
| 24333 | 316 | unfolding word_le_def by auto | 
| 317 | ||
| 24465 | 318 | lemma word_order_refl: "z <= (z :: 'a :: len0 word)" | 
| 24333 | 319 | unfolding word_le_def by auto | 
| 320 | ||
| 24465 | 321 | lemma word_order_antisym: "x <= y ==> y <= x ==> x = (y :: 'a :: len0 word)" | 
| 24333 | 322 | unfolding word_le_def by (auto intro!: word_uint.Rep_eqD) | 
| 323 | ||
| 324 | lemma word_order_linear: | |
| 24465 | 325 | "y <= x | x <= (y :: 'a :: len0 word)" | 
| 24333 | 326 | unfolding word_le_def by auto | 
| 327 | ||
| 328 | lemma word_zero_le [simp] : | |
| 24465 | 329 | "0 <= (y :: 'a :: len0 word)" | 
| 24333 | 330 | unfolding word_le_def by auto | 
| 24465 | 331 | |
| 332 | instance word :: (len0) semigroup_add | |
| 333 | by intro_classes (simp add: word_add_assoc) | |
| 24333 | 334 | |
| 24465 | 335 | instance word :: (len0) linorder | 
| 24333 | 336 | by intro_classes (auto simp: word_less_def word_le_def) | 
| 337 | ||
| 24465 | 338 | instance word :: (len0) ring | 
| 339 | by intro_classes | |
| 340 | (auto simp: word_arith_eqs word_diff_minus | |
| 341 | word_diff_self [unfolded word_diff_minus]) | |
| 342 | ||
| 24333 | 343 | lemma word_m1_ge [simp] : "word_pred 0 >= y" | 
| 344 | unfolding word_le_def | |
| 345 | by (simp only : word_pred_0_n1 word_uint.eq_norm m1mod2k) auto | |
| 346 | ||
| 347 | lemmas word_n1_ge [simp] = word_m1_ge [simplified word_sp_01] | |
| 348 | ||
| 349 | lemmas word_not_simps [simp] = | |
| 350 | word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] | |
| 351 | ||
| 24465 | 352 | lemma word_gt_0: "0 < y = (0 ~= (y :: 'a :: len0 word))" | 
| 24333 | 353 | unfolding word_less_def by auto | 
| 354 | ||
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 355 | lemmas word_gt_0_no [simp] = word_gt_0 [of "number_of y", standard] | 
| 24333 | 356 | |
| 357 | lemma word_sless_alt: "(a <s b) == (sint a < sint b)" | |
| 358 | unfolding word_sle_def word_sless_def | |
| 27682 | 359 | by (auto simp add: less_le) | 
| 24333 | 360 | |
| 361 | lemma word_le_nat_alt: "(a <= b) = (unat a <= unat b)" | |
| 362 | unfolding unat_def word_le_def | |
| 363 | by (rule nat_le_eq_zle [symmetric]) simp | |
| 364 | ||
| 365 | lemma word_less_nat_alt: "(a < b) = (unat a < unat b)" | |
| 366 | unfolding unat_def word_less_alt | |
| 367 | by (rule nat_less_eq_zless [symmetric]) simp | |
| 368 | ||
| 369 | lemma wi_less: | |
| 24465 | 370 | "(word_of_int n < (word_of_int m :: 'a :: len0 word)) = | 
| 371 |     (n mod 2 ^ len_of TYPE('a) < m mod 2 ^ len_of TYPE('a))"
 | |
| 24333 | 372 | unfolding word_less_alt by (simp add: word_uint.eq_norm) | 
| 373 | ||
| 374 | lemma wi_le: | |
| 24465 | 375 | "(word_of_int n <= (word_of_int m :: 'a :: len0 word)) = | 
| 376 |     (n mod 2 ^ len_of TYPE('a) <= m mod 2 ^ len_of TYPE('a))"
 | |
| 24333 | 377 | unfolding word_le_def by (simp add: word_uint.eq_norm) | 
| 378 | ||
| 379 | lemma udvd_nat_alt: "a udvd b = (EX n>=0. unat b = n * unat a)" | |
| 380 | apply (unfold udvd_def) | |
| 381 | apply safe | |
| 382 | apply (simp add: unat_def nat_mult_distrib) | |
| 383 | apply (simp add: uint_nat int_mult) | |
| 384 | apply (rule exI) | |
| 385 | apply safe | |
| 386 | prefer 2 | |
| 387 | apply (erule notE) | |
| 388 | apply (rule refl) | |
| 389 | apply force | |
| 390 | done | |
| 391 | ||
| 392 | lemma udvd_iff_dvd: "x udvd y <-> unat x dvd unat y" | |
| 393 | unfolding dvd_def udvd_nat_alt by force | |
| 394 | ||
| 24465 | 395 | lemmas unat_mono = word_less_nat_alt [THEN iffD1, standard] | 
| 24378 
af83eeb4a702
move udvd, div and mod stuff from WordDefinition to WordArith
 huffman parents: 
24377diff
changeset | 396 | |
| 24465 | 397 | lemma word_zero_neq_one: "0 < len_of TYPE ('a :: len0) ==> (0 :: 'a word) ~= 1";
 | 
| 24333 | 398 | unfolding word_arith_wis | 
| 26514 | 399 | apply (auto simp add: word_ubin.norm_eq_iff [symmetric] gr0_conv_Suc) | 
| 400 | unfolding Bit0_def Bit1_def by simp | |
| 24333 | 401 | |
| 24465 | 402 | lemmas lenw1_zero_neq_one = len_gt_0 [THEN word_zero_neq_one] | 
| 24333 | 403 | |
| 404 | lemma no_no [simp] : "number_of (number_of b) = number_of b" | |
| 405 | by (simp add: number_of_eq) | |
| 406 | ||
| 407 | lemma unat_minus_one: "x ~= 0 ==> unat (x - 1) = unat x - 1" | |
| 408 | apply (unfold unat_def) | |
| 409 | apply (simp only: int_word_uint word_arith_alts rdmods) | |
| 410 | apply (subgoal_tac "uint x >= 1") | |
| 411 | prefer 2 | |
| 412 | apply (drule contrapos_nn) | |
| 413 | apply (erule word_uint.Rep_inverse' [symmetric]) | |
| 414 | apply (insert uint_ge_0 [of x])[1] | |
| 415 | apply arith | |
| 416 | apply (rule box_equals) | |
| 417 | apply (rule nat_diff_distrib) | |
| 418 | prefer 2 | |
| 419 | apply assumption | |
| 420 | apply simp | |
| 421 | apply (subst mod_pos_pos_trivial) | |
| 422 | apply arith | |
| 423 | apply (insert uint_lt2p [of x])[1] | |
| 424 | apply arith | |
| 425 | apply (rule refl) | |
| 426 | apply simp | |
| 427 | done | |
| 428 | ||
| 429 | lemma measure_unat: "p ~= 0 ==> unat (p - 1) < unat p" | |
| 430 | by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) | |
| 431 | ||
| 432 | lemmas uint_add_ge0 [simp] = | |
| 433 | add_nonneg_nonneg [OF uint_ge_0 uint_ge_0, standard] | |
| 434 | lemmas uint_mult_ge0 [simp] = | |
| 435 | mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0, standard] | |
| 436 | ||
| 437 | lemma uint_sub_lt2p [simp]: | |
| 24465 | 438 | "uint (x :: 'a :: len0 word) - uint (y :: 'b :: len0 word) < | 
| 439 |     2 ^ len_of TYPE('a)"
 | |
| 24333 | 440 | using uint_ge_0 [of y] uint_lt2p [of x] by arith | 
| 441 | ||
| 442 | ||
| 24350 | 443 | subsection "Conditions for the addition (etc) of two words to overflow" | 
| 24333 | 444 | |
| 445 | lemma uint_add_lem: | |
| 24465 | 446 |   "(uint x + uint y < 2 ^ len_of TYPE('a)) = 
 | 
| 447 | (uint (x + y :: 'a :: len0 word) = uint x + uint y)" | |
| 24333 | 448 | by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) | 
| 449 | ||
| 450 | lemma uint_mult_lem: | |
| 24465 | 451 |   "(uint x * uint y < 2 ^ len_of TYPE('a)) = 
 | 
| 452 | (uint (x * y :: 'a :: len0 word) = uint x * uint y)" | |
| 24333 | 453 | by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) | 
| 454 | ||
| 455 | lemma uint_sub_lem: | |
| 456 | "(uint x >= uint y) = (uint (x - y) = uint x - uint y)" | |
| 457 | by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) | |
| 458 | ||
| 459 | lemma uint_add_le: "uint (x + y) <= uint x + uint y" | |
| 460 | unfolding uint_word_ariths by (auto simp: mod_add_if_z) | |
| 461 | ||
| 462 | lemma uint_sub_ge: "uint (x - y) >= uint x - uint y" | |
| 463 | unfolding uint_word_ariths by (auto simp: mod_sub_if_z) | |
| 464 | ||
| 465 | lemmas uint_sub_if' = | |
| 466 | trans [OF uint_word_ariths(1) mod_sub_if_z, simplified, standard] | |
| 467 | lemmas uint_plus_if' = | |
| 468 | trans [OF uint_word_ariths(2) mod_add_if_z, simplified, standard] | |
| 469 | ||
| 470 | ||
| 24350 | 471 | subsection {* Definition of uint\_arith *}
 | 
| 24333 | 472 | |
| 473 | lemma word_of_int_inverse: | |
| 24465 | 474 |   "word_of_int r = a ==> 0 <= r ==> r < 2 ^ len_of TYPE('a) ==> 
 | 
| 475 | uint (a::'a::len0 word) = r" | |
| 24333 | 476 | apply (erule word_uint.Abs_inverse' [rotated]) | 
| 477 | apply (simp add: uints_num) | |
| 478 | done | |
| 479 | ||
| 480 | lemma uint_split: | |
| 24465 | 481 | fixes x::"'a::len0 word" | 
| 24333 | 482 | shows "P (uint x) = | 
| 24465 | 483 |          (ALL i. word_of_int i = x & 0 <= i & i < 2^len_of TYPE('a) --> P i)"
 | 
| 24333 | 484 | apply (fold word_int_case_def) | 
| 485 | apply (auto dest!: word_of_int_inverse simp: int_word_uint int_mod_eq' | |
| 486 | split: word_int_split) | |
| 487 | done | |
| 488 | ||
| 489 | lemma uint_split_asm: | |
| 24465 | 490 | fixes x::"'a::len0 word" | 
| 24333 | 491 | shows "P (uint x) = | 
| 24465 | 492 |          (~(EX i. word_of_int i = x & 0 <= i & i < 2^len_of TYPE('a) & ~ P i))"
 | 
| 24333 | 493 | by (auto dest!: word_of_int_inverse | 
| 494 | simp: int_word_uint int_mod_eq' | |
| 495 | split: uint_split) | |
| 496 | ||
| 497 | lemmas uint_splits = uint_split uint_split_asm | |
| 498 | ||
| 499 | lemmas uint_arith_simps = | |
| 500 | word_le_def word_less_alt | |
| 501 | word_uint.Rep_inject [symmetric] | |
| 502 | uint_sub_if' uint_plus_if' | |
| 503 | ||
| 24465 | 504 | (* use this to stop, eg, 2 ^ len_of TYPE (32) being simplified *) | 
| 24333 | 505 | lemma power_False_cong: "False ==> a ^ b = c ^ d" | 
| 506 | by auto | |
| 507 | ||
| 508 | (* uint_arith_tac: reduce to arithmetic on int, try to solve by arith *) | |
| 509 | ML {*
 | |
| 510 | fun uint_arith_ss_of ss = | |
| 511 |   ss addsimps @{thms uint_arith_simps}
 | |
| 512 |      delsimps @{thms word_uint.Rep_inject}
 | |
| 513 |      addsplits @{thms split_if_asm} 
 | |
| 514 |      addcongs @{thms power_False_cong}
 | |
| 515 | ||
| 516 | fun uint_arith_tacs ctxt = | |
| 517 | let fun arith_tac' n t = arith_tac ctxt n t handle COOPER => Seq.empty | |
| 518 | in | |
| 519 | [ CLASET' clarify_tac 1, | |
| 520 | SIMPSET' (full_simp_tac o uint_arith_ss_of) 1, | |
| 521 |       ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms uint_splits} 
 | |
| 522 |                                       addcongs @{thms power_False_cong})),
 | |
| 523 |       rewrite_goals_tac @{thms word_size}, 
 | |
| 524 | ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN | |
| 525 | REPEAT (etac conjE n) THEN | |
| 526 |                          REPEAT (dtac @{thm word_of_int_inverse} n 
 | |
| 527 | THEN atac n | |
| 528 | THEN atac n)), | |
| 529 | TRYALL arith_tac' ] | |
| 530 | end | |
| 531 | ||
| 532 | fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) | |
| 533 | *} | |
| 534 | ||
| 535 | method_setup uint_arith = | |
| 536 | "Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD (uint_arith_tac ctxt 1))" | |
| 537 | "solving word arithmetic via integers and arith" | |
| 538 | ||
| 539 | ||
| 24350 | 540 | subsection "More on overflows and monotonicity" | 
| 24333 | 541 | |
| 542 | lemma no_plus_overflow_uint_size: | |
| 24465 | 543 | "((x :: 'a :: len0 word) <= x + y) = (uint x + uint y < 2 ^ size x)" | 
| 24333 | 544 | unfolding word_size by uint_arith | 
| 545 | ||
| 546 | lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] | |
| 547 | ||
| 24465 | 548 | lemma no_ulen_sub: "((x :: 'a :: len0 word) >= x - y) = (uint y <= uint x)" | 
| 24333 | 549 | by uint_arith | 
| 550 | ||
| 551 | lemma no_olen_add': | |
| 24465 | 552 | fixes x :: "'a::len0 word" | 
| 553 |   shows "(x \<le> y + x) = (uint y + uint x < 2 ^ len_of TYPE('a))"
 | |
| 554 | by (simp add: word_add_ac add_ac no_olen_add) | |
| 24333 | 555 | |
| 556 | lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric], standard] | |
| 557 | ||
| 558 | lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem, standard] | |
| 559 | lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1, standard] | |
| 560 | lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem, standard] | |
| 561 | lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] | |
| 562 | lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] | |
| 563 | lemmas word_sub_le = word_sub_le_iff [THEN iffD2, standard] | |
| 564 | ||
| 565 | lemma word_less_sub1: | |
| 24465 | 566 | "(x :: 'a :: len word) ~= 0 ==> (1 < x) = (0 < x - 1)" | 
| 24333 | 567 | by uint_arith | 
| 568 | ||
| 569 | lemma word_le_sub1: | |
| 24465 | 570 | "(x :: 'a :: len word) ~= 0 ==> (1 <= x) = (0 <= x - 1)" | 
| 24333 | 571 | by uint_arith | 
| 572 | ||
| 573 | lemma sub_wrap_lt: | |
| 24465 | 574 | "((x :: 'a :: len0 word) < x - z) = (x < z)" | 
| 24333 | 575 | by uint_arith | 
| 576 | ||
| 577 | lemma sub_wrap: | |
| 24465 | 578 | "((x :: 'a :: len0 word) <= x - z) = (z = 0 | x < z)" | 
| 24333 | 579 | by uint_arith | 
| 580 | ||
| 581 | lemma plus_minus_not_NULL_ab: | |
| 24465 | 582 | "(x :: 'a :: len0 word) <= ab - c ==> c <= ab ==> c ~= 0 ==> x + c ~= 0" | 
| 24333 | 583 | by uint_arith | 
| 584 | ||
| 585 | lemma plus_minus_no_overflow_ab: | |
| 24465 | 586 | "(x :: 'a :: len0 word) <= ab - c ==> c <= ab ==> x <= x + c" | 
| 24333 | 587 | by uint_arith | 
| 588 | ||
| 589 | lemma le_minus': | |
| 24465 | 590 | "(a :: 'a :: len0 word) + c <= b ==> a <= a + c ==> c <= b - a" | 
| 24333 | 591 | by uint_arith | 
| 592 | ||
| 593 | lemma le_plus': | |
| 24465 | 594 | "(a :: 'a :: len0 word) <= b ==> c <= b - a ==> a + c <= b" | 
| 24333 | 595 | by uint_arith | 
| 596 | ||
| 597 | lemmas le_plus = le_plus' [rotated] | |
| 598 | ||
| 599 | lemmas le_minus = leD [THEN thin_rl, THEN le_minus', standard] | |
| 600 | ||
| 601 | lemma word_plus_mono_right: | |
| 24465 | 602 | "(y :: 'a :: len0 word) <= z ==> x <= x + z ==> x + y <= x + z" | 
| 24333 | 603 | by uint_arith | 
| 604 | ||
| 605 | lemma word_less_minus_cancel: | |
| 24465 | 606 | "y - x < z - x ==> x <= z ==> (y :: 'a :: len0 word) < z" | 
| 24333 | 607 | by uint_arith | 
| 608 | ||
| 609 | lemma word_less_minus_mono_left: | |
| 24465 | 610 | "(y :: 'a :: len0 word) < z ==> x <= y ==> y - x < z - x" | 
| 24333 | 611 | by uint_arith | 
| 612 | ||
| 613 | lemma word_less_minus_mono: | |
| 614 | "a < c ==> d < b ==> a - b < a ==> c - d < c | |
| 24465 | 615 | ==> a - b < c - (d::'a::len word)" | 
| 24333 | 616 | by uint_arith | 
| 617 | ||
| 618 | lemma word_le_minus_cancel: | |
| 24465 | 619 | "y - x <= z - x ==> x <= z ==> (y :: 'a :: len0 word) <= z" | 
| 24333 | 620 | by uint_arith | 
| 621 | ||
| 622 | lemma word_le_minus_mono_left: | |
| 24465 | 623 | "(y :: 'a :: len0 word) <= z ==> x <= y ==> y - x <= z - x" | 
| 24333 | 624 | by uint_arith | 
| 625 | ||
| 626 | lemma word_le_minus_mono: | |
| 627 | "a <= c ==> d <= b ==> a - b <= a ==> c - d <= c | |
| 24465 | 628 | ==> a - b <= c - (d::'a::len word)" | 
| 24333 | 629 | by uint_arith | 
| 630 | ||
| 631 | lemma plus_le_left_cancel_wrap: | |
| 24465 | 632 | "(x :: 'a :: len0 word) + y' < x ==> x + y < x ==> (x + y' < x + y) = (y' < y)" | 
| 24333 | 633 | by uint_arith | 
| 634 | ||
| 635 | lemma plus_le_left_cancel_nowrap: | |
| 24465 | 636 | "(x :: 'a :: len0 word) <= x + y' ==> x <= x + y ==> | 
| 24333 | 637 | (x + y' < x + y) = (y' < y)" | 
| 638 | by uint_arith | |
| 639 | ||
| 640 | lemma word_plus_mono_right2: | |
| 24465 | 641 | "(a :: 'a :: len0 word) <= a + b ==> c <= b ==> a <= a + c" | 
| 24333 | 642 | by uint_arith | 
| 643 | ||
| 644 | lemma word_less_add_right: | |
| 24465 | 645 | "(x :: 'a :: len0 word) < y - z ==> z <= y ==> x + z < y" | 
| 24333 | 646 | by uint_arith | 
| 647 | ||
| 648 | lemma word_less_sub_right: | |
| 24465 | 649 | "(x :: 'a :: len0 word) < y + z ==> y <= x ==> x - y < z" | 
| 24333 | 650 | by uint_arith | 
| 651 | ||
| 652 | lemma word_le_plus_either: | |
| 24465 | 653 | "(x :: 'a :: len0 word) <= y | x <= z ==> y <= y + z ==> x <= y + z" | 
| 24333 | 654 | by uint_arith | 
| 655 | ||
| 656 | lemma word_less_nowrapI: | |
| 24465 | 657 | "(x :: 'a :: len0 word) < z - k ==> k <= z ==> 0 < k ==> x < x + k" | 
| 24333 | 658 | by uint_arith | 
| 659 | ||
| 24465 | 660 | lemma inc_le: "(i :: 'a :: len word) < m ==> i + 1 <= m" | 
| 24333 | 661 | by uint_arith | 
| 662 | ||
| 663 | lemma inc_i: | |
| 24465 | 664 | "(1 :: 'a :: len word) <= i ==> i < m ==> 1 <= (i + 1) & i + 1 <= m" | 
| 24333 | 665 | by uint_arith | 
| 666 | ||
| 667 | lemma udvd_incr_lem: | |
| 668 | "up < uq ==> up = ua + n * uint K ==> | |
| 669 | uq = ua + n' * uint K ==> up + uint K <= uq" | |
| 670 | apply clarsimp | |
| 671 | apply (drule less_le_mult) | |
| 672 | apply safe | |
| 673 | done | |
| 674 | ||
| 675 | lemma udvd_incr': | |
| 676 | "p < q ==> uint p = ua + n * uint K ==> | |
| 677 | uint q = ua + n' * uint K ==> p + K <= q" | |
| 678 | apply (unfold word_less_alt word_le_def) | |
| 679 | apply (drule (2) udvd_incr_lem) | |
| 680 | apply (erule uint_add_le [THEN order_trans]) | |
| 681 | done | |
| 682 | ||
| 683 | lemma udvd_decr': | |
| 684 | "p < q ==> uint p = ua + n * uint K ==> | |
| 685 | uint q = ua + n' * uint K ==> p <= q - K" | |
| 686 | apply (unfold word_less_alt word_le_def) | |
| 687 | apply (drule (2) udvd_incr_lem) | |
| 688 | apply (drule le_diff_eq [THEN iffD2]) | |
| 689 | apply (erule order_trans) | |
| 690 | apply (rule uint_sub_ge) | |
| 691 | done | |
| 692 | ||
| 693 | lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, simplified] | |
| 694 | lemmas udvd_incr0 = udvd_incr' [where ua=0, simplified] | |
| 695 | lemmas udvd_decr0 = udvd_decr' [where ua=0, simplified] | |
| 696 | ||
| 697 | lemma udvd_minus_le': | |
| 698 | "xy < k ==> z udvd xy ==> z udvd k ==> xy <= k - z" | |
| 699 | apply (unfold udvd_def) | |
| 700 | apply clarify | |
| 701 | apply (erule (2) udvd_decr0) | |
| 702 | done | |
| 703 | ||
| 704 | lemma udvd_incr2_K: | |
| 705 | "p < a + s ==> a <= a + s ==> K udvd s ==> K udvd p - a ==> a <= p ==> | |
| 706 | 0 < K ==> p <= p + K & p + K <= a + s" | |
| 707 | apply (unfold udvd_def) | |
| 708 | apply clarify | |
| 709 | apply (simp add: uint_arith_simps split: split_if_asm) | |
| 710 | prefer 2 | |
| 711 | apply (insert uint_range' [of s])[1] | |
| 712 | apply arith | |
| 713 | apply (drule add_commute [THEN xtr1]) | |
| 714 | apply (simp add: diff_less_eq [symmetric]) | |
| 715 | apply (drule less_le_mult) | |
| 716 | apply arith | |
| 717 | apply simp | |
| 718 | done | |
| 719 | ||
| 24465 | 720 | (* links with rbl operations *) | 
| 721 | lemma word_succ_rbl: | |
| 722 | "to_bl w = bl ==> to_bl (word_succ w) = (rev (rbl_succ (rev bl)))" | |
| 723 | apply (unfold word_succ_def) | |
| 724 | apply clarify | |
| 725 | apply (simp add: to_bl_of_bin) | |
| 726 | apply (simp add: to_bl_def rbl_succ) | |
| 727 | done | |
| 728 | ||
| 729 | lemma word_pred_rbl: | |
| 730 | "to_bl w = bl ==> to_bl (word_pred w) = (rev (rbl_pred (rev bl)))" | |
| 731 | apply (unfold word_pred_def) | |
| 732 | apply clarify | |
| 733 | apply (simp add: to_bl_of_bin) | |
| 734 | apply (simp add: to_bl_def rbl_pred) | |
| 735 | done | |
| 736 | ||
| 737 | lemma word_add_rbl: | |
| 738 | "to_bl v = vbl ==> to_bl w = wbl ==> | |
| 739 | to_bl (v + w) = (rev (rbl_add (rev vbl) (rev wbl)))" | |
| 740 | apply (unfold word_add_def) | |
| 741 | apply clarify | |
| 742 | apply (simp add: to_bl_of_bin) | |
| 743 | apply (simp add: to_bl_def rbl_add) | |
| 744 | done | |
| 745 | ||
| 746 | lemma word_mult_rbl: | |
| 747 | "to_bl v = vbl ==> to_bl w = wbl ==> | |
| 748 | to_bl (v * w) = (rev (rbl_mult (rev vbl) (rev wbl)))" | |
| 749 | apply (unfold word_mult_def) | |
| 750 | apply clarify | |
| 751 | apply (simp add: to_bl_of_bin) | |
| 752 | apply (simp add: to_bl_def rbl_mult) | |
| 753 | done | |
| 754 | ||
| 755 | lemma rtb_rbl_ariths: | |
| 756 | "rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_succ w)) = rbl_succ ys" | |
| 757 | ||
| 758 | "rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_pred w)) = rbl_pred ys" | |
| 759 | ||
| 760 | "[| rev (to_bl v) = ys; rev (to_bl w) = xs |] | |
| 761 | ==> rev (to_bl (v * w)) = rbl_mult ys xs" | |
| 762 | ||
| 763 | "[| rev (to_bl v) = ys; rev (to_bl w) = xs |] | |
| 764 | ==> rev (to_bl (v + w)) = rbl_add ys xs" | |
| 765 | by (auto simp: rev_swap [symmetric] word_succ_rbl | |
| 766 | word_pred_rbl word_mult_rbl word_add_rbl) | |
| 767 | ||
| 768 | ||
| 24350 | 769 | subsection "Arithmetic type class instantiations" | 
| 24333 | 770 | |
| 24465 | 771 | instance word :: (len0) comm_monoid_add .. | 
| 772 | ||
| 773 | instance word :: (len0) comm_monoid_mult | |
| 774 | apply (intro_classes) | |
| 775 | apply (simp add: word_mult_commute) | |
| 776 | apply (simp add: word_mult_1) | |
| 777 | done | |
| 778 | ||
| 779 | instance word :: (len0) comm_semiring | |
| 780 | by (intro_classes) (simp add : word_left_distrib) | |
| 781 | ||
| 782 | instance word :: (len0) ab_group_add .. | |
| 783 | ||
| 784 | instance word :: (len0) comm_ring .. | |
| 785 | ||
| 786 | instance word :: (len) comm_semiring_1 | |
| 787 | by (intro_classes) (simp add: lenw1_zero_neq_one) | |
| 788 | ||
| 789 | instance word :: (len) comm_ring_1 .. | |
| 790 | ||
| 791 | instance word :: (len0) comm_semiring_0 .. | |
| 792 | ||
| 793 | instance word :: (len0) order .. | |
| 794 | ||
| 795 | instance word :: (len) recpower | |
| 25762 | 796 | by (intro_classes) simp_all | 
| 24465 | 797 | |
| 24333 | 798 | (* note that iszero_def is only for class comm_semiring_1_cancel, | 
| 24465 | 799 | which requires word length >= 1, ie 'a :: len word *) | 
| 24333 | 800 | lemma zero_bintrunc: | 
| 24465 | 801 | "iszero (number_of x :: 'a :: len word) = | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 802 |     (bintrunc (len_of TYPE('a)) x = Int.Pls)"
 | 
| 24333 | 803 | apply (unfold iszero_def word_0_wi word_no_wi) | 
| 804 | apply (rule word_ubin.norm_eq_iff [symmetric, THEN trans]) | |
| 805 | apply (simp add : Pls_def [symmetric]) | |
| 806 | done | |
| 807 | ||
| 808 | lemmas word_le_0_iff [simp] = | |
| 809 | word_zero_le [THEN leD, THEN linorder_antisym_conv1] | |
| 810 | ||
| 811 | lemma word_of_nat: "of_nat n = word_of_int (int n)" | |
| 812 | by (induct n) (auto simp add : word_of_int_hom_syms) | |
| 813 | ||
| 814 | lemma word_of_int: "of_int = word_of_int" | |
| 815 | apply (rule ext) | |
| 24465 | 816 | apply (unfold of_int_def) | 
| 817 | apply (rule contentsI) | |
| 818 | apply safe | |
| 819 | apply (simp_all add: word_of_nat word_of_int_homs) | |
| 820 | defer | |
| 821 | apply (rule Rep_Integ_ne [THEN nonemptyE]) | |
| 822 | apply (rule bexI) | |
| 823 | prefer 2 | |
| 824 | apply assumption | |
| 825 | apply (auto simp add: RI_eq_diff) | |
| 24333 | 826 | done | 
| 827 | ||
| 828 | lemma word_of_int_nat: | |
| 829 | "0 <= x ==> word_of_int x = of_nat (nat x)" | |
| 830 | by (simp add: of_nat_nat word_of_int) | |
| 831 | ||
| 832 | lemma word_number_of_eq: | |
| 24465 | 833 | "number_of w = (of_int w :: 'a :: len word)" | 
| 24333 | 834 | unfolding word_number_of_def word_of_int by auto | 
| 835 | ||
| 24465 | 836 | instance word :: (len) number_ring | 
| 24333 | 837 | by (intro_classes) (simp add : word_number_of_eq) | 
| 838 | ||
| 839 | lemma iszero_word_no [simp] : | |
| 24465 | 840 | "iszero (number_of bin :: 'a :: len word) = | 
| 841 |     iszero (number_of (bintrunc (len_of TYPE('a)) bin) :: int)"
 | |
| 24368 | 842 | apply (simp add: zero_bintrunc number_of_is_id) | 
| 24333 | 843 | apply (unfold iszero_def Pls_def) | 
| 844 | apply (rule refl) | |
| 845 | done | |
| 846 | ||
| 847 | ||
| 24350 | 848 | subsection "Word and nat" | 
| 24333 | 849 | |
| 850 | lemma td_ext_unat': | |
| 24465 | 851 |   "n = len_of TYPE ('a :: len) ==> 
 | 
| 24333 | 852 | td_ext (unat :: 'a word => nat) of_nat | 
| 853 | (unats n) (%i. i mod 2 ^ n)" | |
| 854 | apply (unfold td_ext_def' unat_def word_of_nat unats_uints) | |
| 855 | apply (auto intro!: imageI simp add : word_of_int_hom_syms) | |
| 856 | apply (erule word_uint.Abs_inverse [THEN arg_cong]) | |
| 857 | apply (simp add: int_word_uint nat_mod_distrib nat_power_eq) | |
| 858 | done | |
| 859 | ||
| 860 | lemmas td_ext_unat = refl [THEN td_ext_unat'] | |
| 861 | lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm, standard] | |
| 862 | ||
| 863 | interpretation word_unat: | |
| 24465 | 864 | td_ext ["unat::'a::len word => nat" | 
| 24333 | 865 | of_nat | 
| 24465 | 866 |           "unats (len_of TYPE('a::len))"
 | 
| 867 |           "%i. i mod 2 ^ len_of TYPE('a::len)"]
 | |
| 24333 | 868 | by (rule td_ext_unat) | 
| 869 | ||
| 870 | lemmas td_unat = word_unat.td_thm | |
| 871 | ||
| 872 | lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] | |
| 873 | ||
| 24465 | 874 | lemma unat_le: "y <= unat (z :: 'a :: len word) ==> y : unats (len_of TYPE ('a))"
 | 
| 24333 | 875 | apply (unfold unats_def) | 
| 876 | apply clarsimp | |
| 877 | apply (rule xtrans, rule unat_lt2p, assumption) | |
| 878 | done | |
| 879 | ||
| 880 | lemma word_nchotomy: | |
| 24465 | 881 |   "ALL w. EX n. (w :: 'a :: len word) = of_nat n & n < 2 ^ len_of TYPE ('a)"
 | 
| 24333 | 882 | apply (rule allI) | 
| 883 | apply (rule word_unat.Abs_cases) | |
| 884 | apply (unfold unats_def) | |
| 885 | apply auto | |
| 886 | done | |
| 887 | ||
| 888 | lemma of_nat_eq: | |
| 24465 | 889 | fixes w :: "'a::len word" | 
| 890 |   shows "(of_nat n = w) = (\<exists>q. n = unat w + q * 2 ^ len_of TYPE('a))"
 | |
| 24333 | 891 | apply (rule trans) | 
| 892 | apply (rule word_unat.inverse_norm) | |
| 893 | apply (rule iffI) | |
| 894 | apply (rule mod_eqD) | |
| 895 | apply simp | |
| 896 | apply clarsimp | |
| 897 | done | |
| 898 | ||
| 899 | lemma of_nat_eq_size: | |
| 900 | "(of_nat n = w) = (EX q. n = unat w + q * 2 ^ size w)" | |
| 901 | unfolding word_size by (rule of_nat_eq) | |
| 902 | ||
| 903 | lemma of_nat_0: | |
| 24465 | 904 |   "(of_nat m = (0::'a::len word)) = (\<exists>q. m = q * 2 ^ len_of TYPE('a))"
 | 
| 24333 | 905 | by (simp add: of_nat_eq) | 
| 906 | ||
| 907 | lemmas of_nat_2p = mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]] | |
| 908 | ||
| 909 | lemma of_nat_gt_0: "of_nat k ~= 0 ==> 0 < k" | |
| 910 | by (cases k) auto | |
| 911 | ||
| 912 | lemma of_nat_neq_0: | |
| 24465 | 913 |   "0 < k ==> k < 2 ^ len_of TYPE ('a :: len) ==> of_nat k ~= (0 :: 'a word)"
 | 
| 24333 | 914 | by (clarsimp simp add : of_nat_0) | 
| 915 | ||
| 916 | lemma Abs_fnat_hom_add: | |
| 917 | "of_nat a + of_nat b = of_nat (a + b)" | |
| 918 | by simp | |
| 919 | ||
| 920 | lemma Abs_fnat_hom_mult: | |
| 24465 | 921 | "of_nat a * of_nat b = (of_nat (a * b) :: 'a :: len word)" | 
| 24333 | 922 | by (simp add: word_of_nat word_of_int_mult_hom zmult_int) | 
| 923 | ||
| 924 | lemma Abs_fnat_hom_Suc: | |
| 925 | "word_succ (of_nat a) = of_nat (Suc a)" | |
| 926 | by (simp add: word_of_nat word_of_int_succ_hom add_ac) | |
| 927 | ||
| 24465 | 928 | lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" | 
| 24333 | 929 | by (simp add: word_of_nat word_0_wi) | 
| 930 | ||
| 24465 | 931 | lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" | 
| 24333 | 932 | by (simp add: word_of_nat word_1_wi) | 
| 933 | ||
| 934 | lemmas Abs_fnat_homs = | |
| 935 | Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc | |
| 936 | Abs_fnat_hom_0 Abs_fnat_hom_1 | |
| 937 | ||
| 938 | lemma word_arith_nat_add: | |
| 939 | "a + b = of_nat (unat a + unat b)" | |
| 940 | by simp | |
| 941 | ||
| 942 | lemma word_arith_nat_mult: | |
| 943 | "a * b = of_nat (unat a * unat b)" | |
| 944 | by (simp add: Abs_fnat_hom_mult [symmetric]) | |
| 945 | ||
| 946 | lemma word_arith_nat_Suc: | |
| 947 | "word_succ a = of_nat (Suc (unat a))" | |
| 948 | by (subst Abs_fnat_hom_Suc [symmetric]) simp | |
| 949 | ||
| 950 | lemma word_arith_nat_div: | |
| 951 | "a div b = of_nat (unat a div unat b)" | |
| 952 | by (simp add: word_div_def word_of_nat zdiv_int uint_nat) | |
| 953 | ||
| 954 | lemma word_arith_nat_mod: | |
| 955 | "a mod b = of_nat (unat a mod unat b)" | |
| 956 | by (simp add: word_mod_def word_of_nat zmod_int uint_nat) | |
| 957 | ||
| 958 | lemmas word_arith_nat_defs = | |
| 959 | word_arith_nat_add word_arith_nat_mult | |
| 960 | word_arith_nat_Suc Abs_fnat_hom_0 | |
| 961 | Abs_fnat_hom_1 word_arith_nat_div | |
| 962 | word_arith_nat_mod | |
| 963 | ||
| 964 | lemmas unat_cong = arg_cong [where f = "unat"] | |
| 965 | ||
| 966 | lemmas unat_word_ariths = word_arith_nat_defs | |
| 967 | [THEN trans [OF unat_cong unat_of_nat], standard] | |
| 968 | ||
| 969 | lemmas word_sub_less_iff = word_sub_le_iff | |
| 970 | [simplified linorder_not_less [symmetric], simplified] | |
| 971 | ||
| 972 | lemma unat_add_lem: | |
| 24465 | 973 |   "(unat x + unat y < 2 ^ len_of TYPE('a)) = 
 | 
| 974 | (unat (x + y :: 'a :: len word) = unat x + unat y)" | |
| 24333 | 975 | unfolding unat_word_ariths | 
| 976 | by (auto intro!: trans [OF _ nat_mod_lem]) | |
| 977 | ||
| 978 | lemma unat_mult_lem: | |
| 24465 | 979 |   "(unat x * unat y < 2 ^ len_of TYPE('a)) = 
 | 
| 980 | (unat (x * y :: 'a :: len word) = unat x * unat y)" | |
| 24333 | 981 | unfolding unat_word_ariths | 
| 982 | by (auto intro!: trans [OF _ nat_mod_lem]) | |
| 983 | ||
| 984 | lemmas unat_plus_if' = | |
| 985 | trans [OF unat_word_ariths(1) mod_nat_add, simplified, standard] | |
| 986 | ||
| 987 | lemma le_no_overflow: | |
| 24465 | 988 | "x <= b ==> a <= a + b ==> x <= a + (b :: 'a :: len0 word)" | 
| 24333 | 989 | apply (erule order_trans) | 
| 990 | apply (erule olen_add_eqv [THEN iffD1]) | |
| 991 | done | |
| 992 | ||
| 993 | lemmas un_ui_le = trans | |
| 994 | [OF word_le_nat_alt [symmetric] | |
| 25762 | 995 | word_le_def, | 
| 24333 | 996 | standard] | 
| 997 | ||
| 998 | lemma unat_sub_if_size: | |
| 999 | "unat (x - y) = (if unat y <= unat x | |
| 1000 | then unat x - unat y | |
| 1001 | else unat x + 2 ^ size x - unat y)" | |
| 1002 | apply (unfold word_size) | |
| 1003 | apply (simp add: un_ui_le) | |
| 1004 | apply (auto simp add: unat_def uint_sub_if') | |
| 1005 | apply (rule nat_diff_distrib) | |
| 1006 | prefer 3 | |
| 1007 | apply (simp add: group_simps) | |
| 1008 | apply (rule nat_diff_distrib [THEN trans]) | |
| 1009 | prefer 3 | |
| 1010 | apply (subst nat_add_distrib) | |
| 1011 | prefer 3 | |
| 1012 | apply (simp add: nat_power_eq) | |
| 1013 | apply auto | |
| 1014 | apply uint_arith | |
| 1015 | done | |
| 1016 | ||
| 1017 | lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] | |
| 1018 | ||
| 24465 | 1019 | lemma unat_div: "unat ((x :: 'a :: len word) div y) = unat x div unat y" | 
| 24333 | 1020 | apply (simp add : unat_word_ariths) | 
| 1021 | apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) | |
| 1022 | apply (rule div_le_dividend) | |
| 1023 | done | |
| 1024 | ||
| 24465 | 1025 | lemma unat_mod: "unat ((x :: 'a :: len word) mod y) = unat x mod unat y" | 
| 24333 | 1026 | apply (clarsimp simp add : unat_word_ariths) | 
| 1027 | apply (cases "unat y") | |
| 1028 | prefer 2 | |
| 1029 | apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) | |
| 1030 | apply (rule mod_le_divisor) | |
| 1031 | apply auto | |
| 1032 | done | |
| 1033 | ||
| 24465 | 1034 | lemma uint_div: "uint ((x :: 'a :: len word) div y) = uint x div uint y" | 
| 24333 | 1035 | unfolding uint_nat by (simp add : unat_div zdiv_int) | 
| 1036 | ||
| 24465 | 1037 | lemma uint_mod: "uint ((x :: 'a :: len word) mod y) = uint x mod uint y" | 
| 24333 | 1038 | unfolding uint_nat by (simp add : unat_mod zmod_int) | 
| 1039 | ||
| 1040 | ||
| 24350 | 1041 | subsection {* Definition of unat\_arith tactic *}
 | 
| 24333 | 1042 | |
| 1043 | lemma unat_split: | |
| 24465 | 1044 | fixes x::"'a::len word" | 
| 24333 | 1045 | shows "P (unat x) = | 
| 24465 | 1046 |          (ALL n. of_nat n = x & n < 2^len_of TYPE('a) --> P n)"
 | 
| 24333 | 1047 | by (auto simp: unat_of_nat) | 
| 1048 | ||
| 1049 | lemma unat_split_asm: | |
| 24465 | 1050 | fixes x::"'a::len word" | 
| 24333 | 1051 | shows "P (unat x) = | 
| 24465 | 1052 |          (~(EX n. of_nat n = x & n < 2^len_of TYPE('a) & ~ P n))"
 | 
| 24333 | 1053 | by (auto simp: unat_of_nat) | 
| 1054 | ||
| 1055 | lemmas of_nat_inverse = | |
| 1056 | word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] | |
| 1057 | ||
| 1058 | lemmas unat_splits = unat_split unat_split_asm | |
| 1059 | ||
| 1060 | lemmas unat_arith_simps = | |
| 1061 | word_le_nat_alt word_less_nat_alt | |
| 1062 | word_unat.Rep_inject [symmetric] | |
| 1063 | unat_sub_if' unat_plus_if' unat_div unat_mod | |
| 1064 | ||
| 1065 | (* unat_arith_tac: tactic to reduce word arithmetic to nat, | |
| 1066 | try to solve via arith *) | |
| 1067 | ML {*
 | |
| 1068 | fun unat_arith_ss_of ss = | |
| 1069 |   ss addsimps @{thms unat_arith_simps}
 | |
| 1070 |      delsimps @{thms word_unat.Rep_inject}
 | |
| 1071 |      addsplits @{thms split_if_asm}
 | |
| 1072 |      addcongs @{thms power_False_cong}
 | |
| 1073 | ||
| 1074 | fun unat_arith_tacs ctxt = | |
| 1075 | let fun arith_tac' n t = arith_tac ctxt n t handle COOPER => Seq.empty | |
| 1076 | in | |
| 1077 | [ CLASET' clarify_tac 1, | |
| 1078 | SIMPSET' (full_simp_tac o unat_arith_ss_of) 1, | |
| 1079 |       ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms unat_splits} 
 | |
| 1080 |                                        addcongs @{thms power_False_cong})),
 | |
| 1081 |       rewrite_goals_tac @{thms word_size}, 
 | |
| 1082 | ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN | |
| 1083 | REPEAT (etac conjE n) THEN | |
| 1084 |                          REPEAT (dtac @{thm of_nat_inverse} n THEN atac n)),
 | |
| 1085 | TRYALL arith_tac' ] | |
| 1086 | end | |
| 1087 | ||
| 1088 | fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) | |
| 1089 | *} | |
| 1090 | ||
| 1091 | method_setup unat_arith = | |
| 1092 | "Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD (unat_arith_tac ctxt 1))" | |
| 1093 | "solving word arithmetic via natural numbers and arith" | |
| 1094 | ||
| 1095 | lemma no_plus_overflow_unat_size: | |
| 24465 | 1096 | "((x :: 'a :: len word) <= x + y) = (unat x + unat y < 2 ^ size x)" | 
| 24333 | 1097 | unfolding word_size by unat_arith | 
| 1098 | ||
| 24465 | 1099 | lemma unat_sub: "b <= a ==> unat (a - b) = unat a - unat (b :: 'a :: len word)" | 
| 24333 | 1100 | by unat_arith | 
| 1101 | ||
| 1102 | lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] | |
| 1103 | ||
| 1104 | lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem, standard] | |
| 1105 | ||
| 1106 | lemma word_div_mult: | |
| 24465 | 1107 |   "(0 :: 'a :: len word) < y ==> unat x * unat y < 2 ^ len_of TYPE('a) ==> 
 | 
| 24333 | 1108 | x * y div y = x" | 
| 1109 | apply unat_arith | |
| 1110 | apply clarsimp | |
| 1111 | apply (subst unat_mult_lem [THEN iffD1]) | |
| 1112 | apply auto | |
| 1113 | done | |
| 1114 | ||
| 24465 | 1115 | lemma div_lt': "(i :: 'a :: len word) <= k div x ==> | 
| 1116 |     unat i * unat x < 2 ^ len_of TYPE('a)"
 | |
| 24333 | 1117 | apply unat_arith | 
| 1118 | apply clarsimp | |
| 1119 | apply (drule mult_le_mono1) | |
| 1120 | apply (erule order_le_less_trans) | |
| 1121 | apply (rule xtr7 [OF unat_lt2p div_mult_le]) | |
| 1122 | done | |
| 1123 | ||
| 1124 | lemmas div_lt'' = order_less_imp_le [THEN div_lt'] | |
| 1125 | ||
| 24465 | 1126 | lemma div_lt_mult: "(i :: 'a :: len word) < k div x ==> 0 < x ==> i * x < k" | 
| 24333 | 1127 | apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) | 
| 1128 | apply (simp add: unat_arith_simps) | |
| 1129 | apply (drule (1) mult_less_mono1) | |
| 1130 | apply (erule order_less_le_trans) | |
| 1131 | apply (rule div_mult_le) | |
| 1132 | done | |
| 1133 | ||
| 1134 | lemma div_le_mult: | |
| 24465 | 1135 | "(i :: 'a :: len word) <= k div x ==> 0 < x ==> i * x <= k" | 
| 24333 | 1136 | apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) | 
| 1137 | apply (simp add: unat_arith_simps) | |
| 1138 | apply (drule mult_le_mono1) | |
| 1139 | apply (erule order_trans) | |
| 1140 | apply (rule div_mult_le) | |
| 1141 | done | |
| 1142 | ||
| 1143 | lemma div_lt_uint': | |
| 24465 | 1144 |   "(i :: 'a :: len word) <= k div x ==> uint i * uint x < 2 ^ len_of TYPE('a)"
 | 
| 24333 | 1145 | apply (unfold uint_nat) | 
| 1146 | apply (drule div_lt') | |
| 1147 | apply (simp add: zmult_int zless_nat_eq_int_zless [symmetric] | |
| 1148 | nat_power_eq) | |
| 1149 | done | |
| 1150 | ||
| 1151 | lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] | |
| 1152 | ||
| 1153 | lemma word_le_exists': | |
| 24465 | 1154 | "(x :: 'a :: len0 word) <= y ==> | 
| 1155 |     (EX z. y = x + z & uint x + uint z < 2 ^ len_of TYPE('a))"
 | |
| 24333 | 1156 | apply (rule exI) | 
| 1157 | apply (rule conjI) | |
| 1158 | apply (rule zadd_diff_inverse) | |
| 1159 | apply uint_arith | |
| 1160 | done | |
| 1161 | ||
| 1162 | lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] | |
| 1163 | ||
| 1164 | lemmas plus_minus_no_overflow = | |
| 1165 | order_less_imp_le [THEN plus_minus_no_overflow_ab] | |
| 1166 | ||
| 1167 | lemmas mcs = word_less_minus_cancel word_less_minus_mono_left | |
| 1168 | word_le_minus_cancel word_le_minus_mono_left | |
| 1169 | ||
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 1170 | lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel, standard] | 
| 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 1171 | lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel, standard] | 
| 24333 | 1172 | lemmas word_plus_mcs = word_diff_ls | 
| 25350 
a5fcf6d12a53
eliminated illegal schematic variables in where/of;
 wenzelm parents: 
25134diff
changeset | 1173 | [where y = "v + x", unfolded add_diff_cancel, standard] | 
| 24333 | 1174 | |
| 1175 | lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] | |
| 1176 | ||
| 1177 | lemmas thd = refl [THEN [2] split_div_lemma [THEN iffD2], THEN conjunct1] | |
| 1178 | ||
| 1179 | lemma thd1: | |
| 1180 | "a div b * b \<le> (a::nat)" | |
| 1181 | using gt_or_eq_0 [of b] | |
| 1182 | apply (rule disjE) | |
| 1183 | apply (erule xtr4 [OF thd mult_commute]) | |
| 1184 | apply clarsimp | |
| 1185 | done | |
| 1186 | ||
| 1187 | lemmas uno_simps [THEN le_unat_uoi, standard] = | |
| 1188 | mod_le_divisor div_le_dividend thd1 | |
| 1189 | ||
| 1190 | lemma word_mod_div_equality: | |
| 24465 | 1191 | "(n div b) * b + (n mod b) = (n :: 'a :: len word)" | 
| 24333 | 1192 | apply (unfold word_less_nat_alt word_arith_nat_defs) | 
| 1193 | apply (cut_tac y="unat b" in gt_or_eq_0) | |
| 1194 | apply (erule disjE) | |
| 1195 | apply (simp add: mod_div_equality uno_simps) | |
| 1196 | apply simp | |
| 1197 | done | |
| 1198 | ||
| 24465 | 1199 | lemma word_div_mult_le: "a div b * b <= (a::'a::len word)" | 
| 24333 | 1200 | apply (unfold word_le_nat_alt word_arith_nat_defs) | 
| 1201 | apply (cut_tac y="unat b" in gt_or_eq_0) | |
| 1202 | apply (erule disjE) | |
| 1203 | apply (simp add: div_mult_le uno_simps) | |
| 1204 | apply simp | |
| 1205 | done | |
| 1206 | ||
| 24465 | 1207 | lemma word_mod_less_divisor: "0 < n ==> m mod n < (n :: 'a :: len word)" | 
| 24333 | 1208 | apply (simp only: word_less_nat_alt word_arith_nat_defs) | 
| 1209 | apply (clarsimp simp add : uno_simps) | |
| 1210 | done | |
| 1211 | ||
| 1212 | lemma word_of_int_power_hom: | |
| 24465 | 1213 | "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a :: len word)" | 
| 24333 | 1214 | by (induct n) (simp_all add : word_of_int_hom_syms power_Suc) | 
| 1215 | ||
| 1216 | lemma word_arith_power_alt: | |
| 24465 | 1217 | "a ^ n = (word_of_int (uint a ^ n) :: 'a :: len word)" | 
| 24333 | 1218 | by (simp add : word_of_int_power_hom [symmetric]) | 
| 1219 | ||
| 24465 | 1220 | lemma of_bl_length_less: | 
| 1221 |   "length x = k ==> k < len_of TYPE('a) ==> (of_bl x :: 'a :: len word) < 2 ^ k"
 | |
| 1222 | apply (unfold of_bl_no [unfolded word_number_of_def] | |
| 1223 | word_less_alt word_number_of_alt) | |
| 1224 | apply safe | |
| 1225 | apply (simp (no_asm) add: word_of_int_power_hom word_uint.eq_norm | |
| 1226 | del: word_of_int_bin) | |
| 1227 | apply (simp add: mod_pos_pos_trivial) | |
| 1228 | apply (subst mod_pos_pos_trivial) | |
| 1229 | apply (rule bl_to_bin_ge0) | |
| 1230 | apply (rule order_less_trans) | |
| 1231 | apply (rule bl_to_bin_lt2p) | |
| 1232 | apply simp | |
| 1233 | apply (rule bl_to_bin_lt2p) | |
| 1234 | done | |
| 1235 | ||
| 24333 | 1236 | |
| 24350 | 1237 | subsection "Cardinality, finiteness of set of words" | 
| 24333 | 1238 | |
| 1239 | lemmas card_lessThan' = card_lessThan [unfolded lessThan_def] | |
| 1240 | ||
| 1241 | lemmas card_eq = word_unat.Abs_inj_on [THEN card_image, | |
| 1242 | unfolded word_unat.image, unfolded unats_def, standard] | |
| 1243 | ||
| 1244 | lemmas card_word = trans [OF card_eq card_lessThan', standard] | |
| 1245 | ||
| 24465 | 1246 | lemma finite_word_UNIV: "finite (UNIV :: 'a :: len word set)" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1247 | apply (rule contrapos_np) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1248 | prefer 2 | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1249 | apply (erule card_infinite) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1250 | apply (simp add: card_word) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1251 | done | 
| 24333 | 1252 | |
| 1253 | lemma card_word_size: | |
| 24465 | 1254 | "card (UNIV :: 'a :: len word set) = (2 ^ size (x :: 'a word))" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1255 | unfolding word_size by (rule card_word) | 
| 24333 | 1256 | |
| 1257 | end | |
| 1258 |