| author | smolkas | 
| Thu, 14 Feb 2013 22:49:22 +0100 | |
| changeset 51129 | 1edc2cc25f19 | 
| parent 50817 | 652731d92061 | 
| child 51143 | 0a2371e7ced3 | 
| permissions | -rw-r--r-- | 
| 47108 | 1 | (* Title: HOL/Num.thy | 
| 2 | Author: Florian Haftmann | |
| 3 | Author: Brian Huffman | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Binary Numerals *}
 | |
| 7 | ||
| 8 | theory Num | |
| 47191 | 9 | imports Datatype | 
| 47108 | 10 | begin | 
| 11 | ||
| 12 | subsection {* The @{text num} type *}
 | |
| 13 | ||
| 14 | datatype num = One | Bit0 num | Bit1 num | |
| 15 | ||
| 16 | text {* Increment function for type @{typ num} *}
 | |
| 17 | ||
| 18 | primrec inc :: "num \<Rightarrow> num" where | |
| 19 | "inc One = Bit0 One" | | |
| 20 | "inc (Bit0 x) = Bit1 x" | | |
| 21 | "inc (Bit1 x) = Bit0 (inc x)" | |
| 22 | ||
| 23 | text {* Converting between type @{typ num} and type @{typ nat} *}
 | |
| 24 | ||
| 25 | primrec nat_of_num :: "num \<Rightarrow> nat" where | |
| 26 | "nat_of_num One = Suc 0" | | |
| 27 | "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | | |
| 28 | "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" | |
| 29 | ||
| 30 | primrec num_of_nat :: "nat \<Rightarrow> num" where | |
| 31 | "num_of_nat 0 = One" | | |
| 32 | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" | |
| 33 | ||
| 34 | lemma nat_of_num_pos: "0 < nat_of_num x" | |
| 35 | by (induct x) simp_all | |
| 36 | ||
| 37 | lemma nat_of_num_neq_0: " nat_of_num x \<noteq> 0" | |
| 38 | by (induct x) simp_all | |
| 39 | ||
| 40 | lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" | |
| 41 | by (induct x) simp_all | |
| 42 | ||
| 43 | lemma num_of_nat_double: | |
| 44 | "0 < n \<Longrightarrow> num_of_nat (n + n) = Bit0 (num_of_nat n)" | |
| 45 | by (induct n) simp_all | |
| 46 | ||
| 47 | text {*
 | |
| 48 |   Type @{typ num} is isomorphic to the strictly positive
 | |
| 49 | natural numbers. | |
| 50 | *} | |
| 51 | ||
| 52 | lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" | |
| 53 | by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) | |
| 54 | ||
| 55 | lemma num_of_nat_inverse: "0 < n \<Longrightarrow> nat_of_num (num_of_nat n) = n" | |
| 56 | by (induct n) (simp_all add: nat_of_num_inc) | |
| 57 | ||
| 58 | lemma num_eq_iff: "x = y \<longleftrightarrow> nat_of_num x = nat_of_num y" | |
| 59 | apply safe | |
| 60 | apply (drule arg_cong [where f=num_of_nat]) | |
| 61 | apply (simp add: nat_of_num_inverse) | |
| 62 | done | |
| 63 | ||
| 64 | lemma num_induct [case_names One inc]: | |
| 65 | fixes P :: "num \<Rightarrow> bool" | |
| 66 | assumes One: "P One" | |
| 67 | and inc: "\<And>x. P x \<Longrightarrow> P (inc x)" | |
| 68 | shows "P x" | |
| 69 | proof - | |
| 70 | obtain n where n: "Suc n = nat_of_num x" | |
| 71 | by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0) | |
| 72 | have "P (num_of_nat (Suc n))" | |
| 73 | proof (induct n) | |
| 74 | case 0 show ?case using One by simp | |
| 75 | next | |
| 76 | case (Suc n) | |
| 77 | then have "P (inc (num_of_nat (Suc n)))" by (rule inc) | |
| 78 | then show "P (num_of_nat (Suc (Suc n)))" by simp | |
| 79 | qed | |
| 80 | with n show "P x" | |
| 81 | by (simp add: nat_of_num_inverse) | |
| 82 | qed | |
| 83 | ||
| 84 | text {*
 | |
| 85 |   From now on, there are two possible models for @{typ num}:
 | |
| 86 |   as positive naturals (rule @{text "num_induct"})
 | |
| 87 |   and as digit representation (rules @{text "num.induct"}, @{text "num.cases"}).
 | |
| 88 | *} | |
| 89 | ||
| 90 | ||
| 91 | subsection {* Numeral operations *}
 | |
| 92 | ||
| 93 | instantiation num :: "{plus,times,linorder}"
 | |
| 94 | begin | |
| 95 | ||
| 96 | definition [code del]: | |
| 97 | "m + n = num_of_nat (nat_of_num m + nat_of_num n)" | |
| 98 | ||
| 99 | definition [code del]: | |
| 100 | "m * n = num_of_nat (nat_of_num m * nat_of_num n)" | |
| 101 | ||
| 102 | definition [code del]: | |
| 103 | "m \<le> n \<longleftrightarrow> nat_of_num m \<le> nat_of_num n" | |
| 104 | ||
| 105 | definition [code del]: | |
| 106 | "m < n \<longleftrightarrow> nat_of_num m < nat_of_num n" | |
| 107 | ||
| 108 | instance | |
| 109 | by (default, auto simp add: less_num_def less_eq_num_def num_eq_iff) | |
| 110 | ||
| 111 | end | |
| 112 | ||
| 113 | lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" | |
| 114 | unfolding plus_num_def | |
| 115 | by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) | |
| 116 | ||
| 117 | lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" | |
| 118 | unfolding times_num_def | |
| 119 | by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) | |
| 120 | ||
| 121 | lemma add_num_simps [simp, code]: | |
| 122 | "One + One = Bit0 One" | |
| 123 | "One + Bit0 n = Bit1 n" | |
| 124 | "One + Bit1 n = Bit0 (n + One)" | |
| 125 | "Bit0 m + One = Bit1 m" | |
| 126 | "Bit0 m + Bit0 n = Bit0 (m + n)" | |
| 127 | "Bit0 m + Bit1 n = Bit1 (m + n)" | |
| 128 | "Bit1 m + One = Bit0 (m + One)" | |
| 129 | "Bit1 m + Bit0 n = Bit1 (m + n)" | |
| 130 | "Bit1 m + Bit1 n = Bit0 (m + n + One)" | |
| 131 | by (simp_all add: num_eq_iff nat_of_num_add) | |
| 132 | ||
| 133 | lemma mult_num_simps [simp, code]: | |
| 134 | "m * One = m" | |
| 135 | "One * n = n" | |
| 136 | "Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))" | |
| 137 | "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" | |
| 138 | "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" | |
| 139 | "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" | |
| 140 | by (simp_all add: num_eq_iff nat_of_num_add | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49690diff
changeset | 141 | nat_of_num_mult distrib_right distrib_left) | 
| 47108 | 142 | |
| 143 | lemma eq_num_simps: | |
| 144 | "One = One \<longleftrightarrow> True" | |
| 145 | "One = Bit0 n \<longleftrightarrow> False" | |
| 146 | "One = Bit1 n \<longleftrightarrow> False" | |
| 147 | "Bit0 m = One \<longleftrightarrow> False" | |
| 148 | "Bit1 m = One \<longleftrightarrow> False" | |
| 149 | "Bit0 m = Bit0 n \<longleftrightarrow> m = n" | |
| 150 | "Bit0 m = Bit1 n \<longleftrightarrow> False" | |
| 151 | "Bit1 m = Bit0 n \<longleftrightarrow> False" | |
| 152 | "Bit1 m = Bit1 n \<longleftrightarrow> m = n" | |
| 153 | by simp_all | |
| 154 | ||
| 155 | lemma le_num_simps [simp, code]: | |
| 156 | "One \<le> n \<longleftrightarrow> True" | |
| 157 | "Bit0 m \<le> One \<longleftrightarrow> False" | |
| 158 | "Bit1 m \<le> One \<longleftrightarrow> False" | |
| 159 | "Bit0 m \<le> Bit0 n \<longleftrightarrow> m \<le> n" | |
| 160 | "Bit0 m \<le> Bit1 n \<longleftrightarrow> m \<le> n" | |
| 161 | "Bit1 m \<le> Bit1 n \<longleftrightarrow> m \<le> n" | |
| 162 | "Bit1 m \<le> Bit0 n \<longleftrightarrow> m < n" | |
| 163 | using nat_of_num_pos [of n] nat_of_num_pos [of m] | |
| 164 | by (auto simp add: less_eq_num_def less_num_def) | |
| 165 | ||
| 166 | lemma less_num_simps [simp, code]: | |
| 167 | "m < One \<longleftrightarrow> False" | |
| 168 | "One < Bit0 n \<longleftrightarrow> True" | |
| 169 | "One < Bit1 n \<longleftrightarrow> True" | |
| 170 | "Bit0 m < Bit0 n \<longleftrightarrow> m < n" | |
| 171 | "Bit0 m < Bit1 n \<longleftrightarrow> m \<le> n" | |
| 172 | "Bit1 m < Bit1 n \<longleftrightarrow> m < n" | |
| 173 | "Bit1 m < Bit0 n \<longleftrightarrow> m < n" | |
| 174 | using nat_of_num_pos [of n] nat_of_num_pos [of m] | |
| 175 | by (auto simp add: less_eq_num_def less_num_def) | |
| 176 | ||
| 177 | text {* Rules using @{text One} and @{text inc} as constructors *}
 | |
| 178 | ||
| 179 | lemma add_One: "x + One = inc x" | |
| 180 | by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) | |
| 181 | ||
| 182 | lemma add_One_commute: "One + n = n + One" | |
| 183 | by (induct n) simp_all | |
| 184 | ||
| 185 | lemma add_inc: "x + inc y = inc (x + y)" | |
| 186 | by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) | |
| 187 | ||
| 188 | lemma mult_inc: "x * inc y = x * y + x" | |
| 189 | by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) | |
| 190 | ||
| 191 | text {* The @{const num_of_nat} conversion *}
 | |
| 192 | ||
| 193 | lemma num_of_nat_One: | |
| 47300 | 194 | "n \<le> 1 \<Longrightarrow> num_of_nat n = One" | 
| 47108 | 195 | by (cases n) simp_all | 
| 196 | ||
| 197 | lemma num_of_nat_plus_distrib: | |
| 198 | "0 < m \<Longrightarrow> 0 < n \<Longrightarrow> num_of_nat (m + n) = num_of_nat m + num_of_nat n" | |
| 199 | by (induct n) (auto simp add: add_One add_One_commute add_inc) | |
| 200 | ||
| 201 | text {* A double-and-decrement function *}
 | |
| 202 | ||
| 203 | primrec BitM :: "num \<Rightarrow> num" where | |
| 204 | "BitM One = One" | | |
| 205 | "BitM (Bit0 n) = Bit1 (BitM n)" | | |
| 206 | "BitM (Bit1 n) = Bit1 (Bit0 n)" | |
| 207 | ||
| 208 | lemma BitM_plus_one: "BitM n + One = Bit0 n" | |
| 209 | by (induct n) simp_all | |
| 210 | ||
| 211 | lemma one_plus_BitM: "One + BitM n = Bit0 n" | |
| 212 | unfolding add_One_commute BitM_plus_one .. | |
| 213 | ||
| 214 | text {* Squaring and exponentiation *}
 | |
| 215 | ||
| 216 | primrec sqr :: "num \<Rightarrow> num" where | |
| 217 | "sqr One = One" | | |
| 218 | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | | |
| 219 | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" | |
| 220 | ||
| 221 | primrec pow :: "num \<Rightarrow> num \<Rightarrow> num" where | |
| 222 | "pow x One = x" | | |
| 223 | "pow x (Bit0 y) = sqr (pow x y)" | | |
| 47191 | 224 | "pow x (Bit1 y) = sqr (pow x y) * x" | 
| 47108 | 225 | |
| 226 | lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" | |
| 227 | by (induct x, simp_all add: algebra_simps nat_of_num_add) | |
| 228 | ||
| 229 | lemma sqr_conv_mult: "sqr x = x * x" | |
| 230 | by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) | |
| 231 | ||
| 232 | ||
| 47211 | 233 | subsection {* Binary numerals *}
 | 
| 47108 | 234 | |
| 235 | text {*
 | |
| 47211 | 236 | We embed binary representations into a generic algebraic | 
| 47108 | 237 |   structure using @{text numeral}.
 | 
| 238 | *} | |
| 239 | ||
| 240 | class numeral = one + semigroup_add | |
| 241 | begin | |
| 242 | ||
| 243 | primrec numeral :: "num \<Rightarrow> 'a" where | |
| 244 | numeral_One: "numeral One = 1" | | |
| 245 | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | | |
| 246 | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" | |
| 247 | ||
| 50817 | 248 | lemma numeral_code [code]: | 
| 249 | "numeral One = 1" | |
| 250 | "numeral (Bit0 n) = (let m = numeral n in m + m)" | |
| 251 | "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" | |
| 252 | by (simp_all add: Let_def) | |
| 253 | ||
| 47108 | 254 | lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" | 
| 255 | apply (induct x) | |
| 256 | apply simp | |
| 257 | apply (simp add: add_assoc [symmetric], simp add: add_assoc) | |
| 258 | apply (simp add: add_assoc [symmetric], simp add: add_assoc) | |
| 259 | done | |
| 260 | ||
| 261 | lemma numeral_inc: "numeral (inc x) = numeral x + 1" | |
| 262 | proof (induct x) | |
| 263 | case (Bit1 x) | |
| 264 | have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" | |
| 265 | by (simp only: one_plus_numeral_commute) | |
| 266 | with Bit1 show ?case | |
| 267 | by (simp add: add_assoc) | |
| 268 | qed simp_all | |
| 269 | ||
| 270 | declare numeral.simps [simp del] | |
| 271 | ||
| 272 | abbreviation "Numeral1 \<equiv> numeral One" | |
| 273 | ||
| 274 | declare numeral_One [code_post] | |
| 275 | ||
| 276 | end | |
| 277 | ||
| 278 | text {* Negative numerals. *}
 | |
| 279 | ||
| 280 | class neg_numeral = numeral + group_add | |
| 281 | begin | |
| 282 | ||
| 283 | definition neg_numeral :: "num \<Rightarrow> 'a" where | |
| 284 | "neg_numeral k = - numeral k" | |
| 285 | ||
| 286 | end | |
| 287 | ||
| 288 | text {* Numeral syntax. *}
 | |
| 289 | ||
| 290 | syntax | |
| 291 |   "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
 | |
| 292 | ||
| 293 | parse_translation {*
 | |
| 294 | let | |
| 295 | fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2) | |
| 296 |      of (0, 1) => Syntax.const @{const_name One}
 | |
| 297 |       | (n, 0) => Syntax.const @{const_name Bit0} $ num_of_int n
 | |
| 298 |       | (n, 1) => Syntax.const @{const_name Bit1} $ num_of_int n
 | |
| 299 | else raise Match; | |
| 300 |   val pos = Syntax.const @{const_name numeral}
 | |
| 301 |   val neg = Syntax.const @{const_name neg_numeral}
 | |
| 302 |   val one = Syntax.const @{const_name Groups.one}
 | |
| 303 |   val zero = Syntax.const @{const_name Groups.zero}
 | |
| 304 |   fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
 | |
| 305 | c $ numeral_tr [t] $ u | |
| 306 | | numeral_tr [Const (num, _)] = | |
| 307 | let | |
| 308 |           val {value, ...} = Lexicon.read_xnum num;
 | |
| 309 | in | |
| 310 | if value = 0 then zero else | |
| 311 | if value > 0 | |
| 312 | then pos $ num_of_int value | |
| 313 | else neg $ num_of_int (~value) | |
| 314 | end | |
| 315 |     | numeral_tr ts = raise TERM ("numeral_tr", ts);
 | |
| 316 | in [("_Numeral", numeral_tr)] end
 | |
| 317 | *} | |
| 318 | ||
| 319 | typed_print_translation (advanced) {*
 | |
| 320 | let | |
| 321 |   fun dest_num (Const (@{const_syntax Bit0}, _) $ n) = 2 * dest_num n
 | |
| 322 |     | dest_num (Const (@{const_syntax Bit1}, _) $ n) = 2 * dest_num n + 1
 | |
| 323 |     | dest_num (Const (@{const_syntax One}, _)) = 1;
 | |
| 324 | fun num_tr' sign ctxt T [n] = | |
| 325 | let | |
| 326 | val k = dest_num n; | |
| 327 |       val t' = Syntax.const @{syntax_const "_Numeral"} $
 | |
| 328 | Syntax.free (sign ^ string_of_int k); | |
| 329 | in | |
| 330 | case T of | |
| 331 |         Type (@{type_name fun}, [_, T']) =>
 | |
| 49690 
a6814de45b69
more explicit show_type_constraint, show_sort_constraint;
 wenzelm parents: 
48891diff
changeset | 332 | if not (Printer.show_type_constraint ctxt) andalso can Term.dest_Type T' then t' | 
| 47108 | 333 |           else Syntax.const @{syntax_const "_constrain"} $ t' $ Syntax_Phases.term_of_typ ctxt T'
 | 
| 334 | | T' => if T' = dummyT then t' else raise Match | |
| 335 | end; | |
| 336 | in [(@{const_syntax numeral}, num_tr' ""),
 | |
| 337 |     (@{const_syntax neg_numeral}, num_tr' "-")] end
 | |
| 338 | *} | |
| 339 | ||
| 48891 | 340 | ML_file "Tools/numeral.ML" | 
| 47228 | 341 | |
| 342 | ||
| 47108 | 343 | subsection {* Class-specific numeral rules *}
 | 
| 344 | ||
| 345 | text {*
 | |
| 346 |   @{const numeral} is a morphism.
 | |
| 347 | *} | |
| 348 | ||
| 349 | subsubsection {* Structures with addition: class @{text numeral} *}
 | |
| 350 | ||
| 351 | context numeral | |
| 352 | begin | |
| 353 | ||
| 354 | lemma numeral_add: "numeral (m + n) = numeral m + numeral n" | |
| 355 | by (induct n rule: num_induct) | |
| 356 | (simp_all only: numeral_One add_One add_inc numeral_inc add_assoc) | |
| 357 | ||
| 358 | lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" | |
| 359 | by (rule numeral_add [symmetric]) | |
| 360 | ||
| 361 | lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" | |
| 362 | using numeral_add [of n One] by (simp add: numeral_One) | |
| 363 | ||
| 364 | lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" | |
| 365 | using numeral_add [of One n] by (simp add: numeral_One) | |
| 366 | ||
| 367 | lemma one_add_one: "1 + 1 = 2" | |
| 368 | using numeral_add [of One One] by (simp add: numeral_One) | |
| 369 | ||
| 370 | lemmas add_numeral_special = | |
| 371 | numeral_plus_one one_plus_numeral one_add_one | |
| 372 | ||
| 373 | end | |
| 374 | ||
| 375 | subsubsection {*
 | |
| 376 |   Structures with negation: class @{text neg_numeral}
 | |
| 377 | *} | |
| 378 | ||
| 379 | context neg_numeral | |
| 380 | begin | |
| 381 | ||
| 382 | text {* Numerals form an abelian subgroup. *}
 | |
| 383 | ||
| 384 | inductive is_num :: "'a \<Rightarrow> bool" where | |
| 385 | "is_num 1" | | |
| 386 | "is_num x \<Longrightarrow> is_num (- x)" | | |
| 387 | "\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> is_num (x + y)" | |
| 388 | ||
| 389 | lemma is_num_numeral: "is_num (numeral k)" | |
| 390 | by (induct k, simp_all add: numeral.simps is_num.intros) | |
| 391 | ||
| 392 | lemma is_num_add_commute: | |
| 393 | "\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> x + y = y + x" | |
| 394 | apply (induct x rule: is_num.induct) | |
| 395 | apply (induct y rule: is_num.induct) | |
| 396 | apply simp | |
| 397 | apply (rule_tac a=x in add_left_imp_eq) | |
| 398 | apply (rule_tac a=x in add_right_imp_eq) | |
| 399 | apply (simp add: add_assoc minus_add_cancel) | |
| 400 | apply (simp add: add_assoc [symmetric], simp add: add_assoc) | |
| 401 | apply (rule_tac a=x in add_left_imp_eq) | |
| 402 | apply (rule_tac a=x in add_right_imp_eq) | |
| 403 | apply (simp add: add_assoc minus_add_cancel add_minus_cancel) | |
| 404 | apply (simp add: add_assoc, simp add: add_assoc [symmetric]) | |
| 405 | done | |
| 406 | ||
| 407 | lemma is_num_add_left_commute: | |
| 408 | "\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> x + (y + z) = y + (x + z)" | |
| 409 | by (simp only: add_assoc [symmetric] is_num_add_commute) | |
| 410 | ||
| 411 | lemmas is_num_normalize = | |
| 412 | add_assoc is_num_add_commute is_num_add_left_commute | |
| 413 | is_num.intros is_num_numeral | |
| 414 | diff_minus minus_add add_minus_cancel minus_add_cancel | |
| 415 | ||
| 416 | definition dbl :: "'a \<Rightarrow> 'a" where "dbl x = x + x" | |
| 417 | definition dbl_inc :: "'a \<Rightarrow> 'a" where "dbl_inc x = x + x + 1" | |
| 418 | definition dbl_dec :: "'a \<Rightarrow> 'a" where "dbl_dec x = x + x - 1" | |
| 419 | ||
| 420 | definition sub :: "num \<Rightarrow> num \<Rightarrow> 'a" where | |
| 421 | "sub k l = numeral k - numeral l" | |
| 422 | ||
| 423 | lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" | |
| 424 | by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) | |
| 425 | ||
| 426 | lemma dbl_simps [simp]: | |
| 427 | "dbl (neg_numeral k) = neg_numeral (Bit0 k)" | |
| 428 | "dbl 0 = 0" | |
| 429 | "dbl 1 = 2" | |
| 430 | "dbl (numeral k) = numeral (Bit0 k)" | |
| 431 | unfolding dbl_def neg_numeral_def numeral.simps | |
| 432 | by (simp_all add: minus_add) | |
| 433 | ||
| 434 | lemma dbl_inc_simps [simp]: | |
| 435 | "dbl_inc (neg_numeral k) = neg_numeral (BitM k)" | |
| 436 | "dbl_inc 0 = 1" | |
| 437 | "dbl_inc 1 = 3" | |
| 438 | "dbl_inc (numeral k) = numeral (Bit1 k)" | |
| 439 | unfolding dbl_inc_def neg_numeral_def numeral.simps numeral_BitM | |
| 440 | by (simp_all add: is_num_normalize) | |
| 441 | ||
| 442 | lemma dbl_dec_simps [simp]: | |
| 443 | "dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)" | |
| 444 | "dbl_dec 0 = -1" | |
| 445 | "dbl_dec 1 = 1" | |
| 446 | "dbl_dec (numeral k) = numeral (BitM k)" | |
| 447 | unfolding dbl_dec_def neg_numeral_def numeral.simps numeral_BitM | |
| 448 | by (simp_all add: is_num_normalize) | |
| 449 | ||
| 450 | lemma sub_num_simps [simp]: | |
| 451 | "sub One One = 0" | |
| 452 | "sub One (Bit0 l) = neg_numeral (BitM l)" | |
| 453 | "sub One (Bit1 l) = neg_numeral (Bit0 l)" | |
| 454 | "sub (Bit0 k) One = numeral (BitM k)" | |
| 455 | "sub (Bit1 k) One = numeral (Bit0 k)" | |
| 456 | "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" | |
| 457 | "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" | |
| 458 | "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" | |
| 459 | "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" | |
| 460 | unfolding dbl_def dbl_dec_def dbl_inc_def sub_def | |
| 461 | unfolding neg_numeral_def numeral.simps numeral_BitM | |
| 462 | by (simp_all add: is_num_normalize) | |
| 463 | ||
| 464 | lemma add_neg_numeral_simps: | |
| 465 | "numeral m + neg_numeral n = sub m n" | |
| 466 | "neg_numeral m + numeral n = sub n m" | |
| 467 | "neg_numeral m + neg_numeral n = neg_numeral (m + n)" | |
| 468 | unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps | |
| 469 | by (simp_all add: is_num_normalize) | |
| 470 | ||
| 471 | lemma add_neg_numeral_special: | |
| 472 | "1 + neg_numeral m = sub One m" | |
| 473 | "neg_numeral m + 1 = sub One m" | |
| 474 | unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps | |
| 475 | by (simp_all add: is_num_normalize) | |
| 476 | ||
| 477 | lemma diff_numeral_simps: | |
| 478 | "numeral m - numeral n = sub m n" | |
| 479 | "numeral m - neg_numeral n = numeral (m + n)" | |
| 480 | "neg_numeral m - numeral n = neg_numeral (m + n)" | |
| 481 | "neg_numeral m - neg_numeral n = sub n m" | |
| 482 | unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps | |
| 483 | by (simp_all add: is_num_normalize) | |
| 484 | ||
| 485 | lemma diff_numeral_special: | |
| 486 | "1 - numeral n = sub One n" | |
| 487 | "1 - neg_numeral n = numeral (One + n)" | |
| 488 | "numeral m - 1 = sub m One" | |
| 489 | "neg_numeral m - 1 = neg_numeral (m + One)" | |
| 490 | unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps | |
| 491 | by (simp_all add: is_num_normalize) | |
| 492 | ||
| 493 | lemma minus_one: "- 1 = -1" | |
| 494 | unfolding neg_numeral_def numeral.simps .. | |
| 495 | ||
| 496 | lemma minus_numeral: "- numeral n = neg_numeral n" | |
| 497 | unfolding neg_numeral_def .. | |
| 498 | ||
| 499 | lemma minus_neg_numeral: "- neg_numeral n = numeral n" | |
| 500 | unfolding neg_numeral_def by simp | |
| 501 | ||
| 502 | lemmas minus_numeral_simps [simp] = | |
| 503 | minus_one minus_numeral minus_neg_numeral | |
| 504 | ||
| 505 | end | |
| 506 | ||
| 507 | subsubsection {*
 | |
| 508 |   Structures with multiplication: class @{text semiring_numeral}
 | |
| 509 | *} | |
| 510 | ||
| 511 | class semiring_numeral = semiring + monoid_mult | |
| 512 | begin | |
| 513 | ||
| 514 | subclass numeral .. | |
| 515 | ||
| 516 | lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" | |
| 517 | apply (induct n rule: num_induct) | |
| 518 | apply (simp add: numeral_One) | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49690diff
changeset | 519 | apply (simp add: mult_inc numeral_inc numeral_add distrib_left) | 
| 47108 | 520 | done | 
| 521 | ||
| 522 | lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" | |
| 523 | by (rule numeral_mult [symmetric]) | |
| 524 | ||
| 525 | end | |
| 526 | ||
| 527 | subsubsection {*
 | |
| 528 |   Structures with a zero: class @{text semiring_1}
 | |
| 529 | *} | |
| 530 | ||
| 531 | context semiring_1 | |
| 532 | begin | |
| 533 | ||
| 534 | subclass semiring_numeral .. | |
| 535 | ||
| 536 | lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" | |
| 537 | by (induct n, | |
| 538 | simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) | |
| 539 | ||
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 540 | lemma mult_2: "2 * z = z + z" | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49690diff
changeset | 541 | unfolding one_add_one [symmetric] distrib_right by simp | 
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 542 | |
| 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 543 | lemma mult_2_right: "z * 2 = z + z" | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49690diff
changeset | 544 | unfolding one_add_one [symmetric] distrib_left by simp | 
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 545 | |
| 47108 | 546 | end | 
| 547 | ||
| 548 | lemma nat_of_num_numeral: "nat_of_num = numeral" | |
| 549 | proof | |
| 550 | fix n | |
| 551 | have "numeral n = nat_of_num n" | |
| 552 | by (induct n) (simp_all add: numeral.simps) | |
| 553 | then show "nat_of_num n = numeral n" by simp | |
| 554 | qed | |
| 555 | ||
| 556 | subsubsection {*
 | |
| 557 |   Equality: class @{text semiring_char_0}
 | |
| 558 | *} | |
| 559 | ||
| 560 | context semiring_char_0 | |
| 561 | begin | |
| 562 | ||
| 563 | lemma numeral_eq_iff: "numeral m = numeral n \<longleftrightarrow> m = n" | |
| 564 | unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] | |
| 565 | of_nat_eq_iff num_eq_iff .. | |
| 566 | ||
| 567 | lemma numeral_eq_one_iff: "numeral n = 1 \<longleftrightarrow> n = One" | |
| 568 | by (rule numeral_eq_iff [of n One, unfolded numeral_One]) | |
| 569 | ||
| 570 | lemma one_eq_numeral_iff: "1 = numeral n \<longleftrightarrow> One = n" | |
| 571 | by (rule numeral_eq_iff [of One n, unfolded numeral_One]) | |
| 572 | ||
| 573 | lemma numeral_neq_zero: "numeral n \<noteq> 0" | |
| 574 | unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] | |
| 575 | by (simp add: nat_of_num_pos) | |
| 576 | ||
| 577 | lemma zero_neq_numeral: "0 \<noteq> numeral n" | |
| 578 | unfolding eq_commute [of 0] by (rule numeral_neq_zero) | |
| 579 | ||
| 580 | lemmas eq_numeral_simps [simp] = | |
| 581 | numeral_eq_iff | |
| 582 | numeral_eq_one_iff | |
| 583 | one_eq_numeral_iff | |
| 584 | numeral_neq_zero | |
| 585 | zero_neq_numeral | |
| 586 | ||
| 587 | end | |
| 588 | ||
| 589 | subsubsection {*
 | |
| 590 |   Comparisons: class @{text linordered_semidom}
 | |
| 591 | *} | |
| 592 | ||
| 593 | text {*  Could be perhaps more general than here. *}
 | |
| 594 | ||
| 595 | context linordered_semidom | |
| 596 | begin | |
| 597 | ||
| 598 | lemma numeral_le_iff: "numeral m \<le> numeral n \<longleftrightarrow> m \<le> n" | |
| 599 | proof - | |
| 600 | have "of_nat (numeral m) \<le> of_nat (numeral n) \<longleftrightarrow> m \<le> n" | |
| 601 | unfolding less_eq_num_def nat_of_num_numeral of_nat_le_iff .. | |
| 602 | then show ?thesis by simp | |
| 603 | qed | |
| 604 | ||
| 605 | lemma one_le_numeral: "1 \<le> numeral n" | |
| 606 | using numeral_le_iff [of One n] by (simp add: numeral_One) | |
| 607 | ||
| 608 | lemma numeral_le_one_iff: "numeral n \<le> 1 \<longleftrightarrow> n \<le> One" | |
| 609 | using numeral_le_iff [of n One] by (simp add: numeral_One) | |
| 610 | ||
| 611 | lemma numeral_less_iff: "numeral m < numeral n \<longleftrightarrow> m < n" | |
| 612 | proof - | |
| 613 | have "of_nat (numeral m) < of_nat (numeral n) \<longleftrightarrow> m < n" | |
| 614 | unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. | |
| 615 | then show ?thesis by simp | |
| 616 | qed | |
| 617 | ||
| 618 | lemma not_numeral_less_one: "\<not> numeral n < 1" | |
| 619 | using numeral_less_iff [of n One] by (simp add: numeral_One) | |
| 620 | ||
| 621 | lemma one_less_numeral_iff: "1 < numeral n \<longleftrightarrow> One < n" | |
| 622 | using numeral_less_iff [of One n] by (simp add: numeral_One) | |
| 623 | ||
| 624 | lemma zero_le_numeral: "0 \<le> numeral n" | |
| 625 | by (induct n) (simp_all add: numeral.simps) | |
| 626 | ||
| 627 | lemma zero_less_numeral: "0 < numeral n" | |
| 628 | by (induct n) (simp_all add: numeral.simps add_pos_pos) | |
| 629 | ||
| 630 | lemma not_numeral_le_zero: "\<not> numeral n \<le> 0" | |
| 631 | by (simp add: not_le zero_less_numeral) | |
| 632 | ||
| 633 | lemma not_numeral_less_zero: "\<not> numeral n < 0" | |
| 634 | by (simp add: not_less zero_le_numeral) | |
| 635 | ||
| 636 | lemmas le_numeral_extra = | |
| 637 | zero_le_one not_one_le_zero | |
| 638 | order_refl [of 0] order_refl [of 1] | |
| 639 | ||
| 640 | lemmas less_numeral_extra = | |
| 641 | zero_less_one not_one_less_zero | |
| 642 | less_irrefl [of 0] less_irrefl [of 1] | |
| 643 | ||
| 644 | lemmas le_numeral_simps [simp] = | |
| 645 | numeral_le_iff | |
| 646 | one_le_numeral | |
| 647 | numeral_le_one_iff | |
| 648 | zero_le_numeral | |
| 649 | not_numeral_le_zero | |
| 650 | ||
| 651 | lemmas less_numeral_simps [simp] = | |
| 652 | numeral_less_iff | |
| 653 | one_less_numeral_iff | |
| 654 | not_numeral_less_one | |
| 655 | zero_less_numeral | |
| 656 | not_numeral_less_zero | |
| 657 | ||
| 658 | end | |
| 659 | ||
| 660 | subsubsection {*
 | |
| 661 |   Multiplication and negation: class @{text ring_1}
 | |
| 662 | *} | |
| 663 | ||
| 664 | context ring_1 | |
| 665 | begin | |
| 666 | ||
| 667 | subclass neg_numeral .. | |
| 668 | ||
| 669 | lemma mult_neg_numeral_simps: | |
| 670 | "neg_numeral m * neg_numeral n = numeral (m * n)" | |
| 671 | "neg_numeral m * numeral n = neg_numeral (m * n)" | |
| 672 | "numeral m * neg_numeral n = neg_numeral (m * n)" | |
| 673 | unfolding neg_numeral_def mult_minus_left mult_minus_right | |
| 674 | by (simp_all only: minus_minus numeral_mult) | |
| 675 | ||
| 676 | lemma mult_minus1 [simp]: "-1 * z = - z" | |
| 677 | unfolding neg_numeral_def numeral.simps mult_minus_left by simp | |
| 678 | ||
| 679 | lemma mult_minus1_right [simp]: "z * -1 = - z" | |
| 680 | unfolding neg_numeral_def numeral.simps mult_minus_right by simp | |
| 681 | ||
| 682 | end | |
| 683 | ||
| 684 | subsubsection {*
 | |
| 685 |   Equality using @{text iszero} for rings with non-zero characteristic
 | |
| 686 | *} | |
| 687 | ||
| 688 | context ring_1 | |
| 689 | begin | |
| 690 | ||
| 691 | definition iszero :: "'a \<Rightarrow> bool" | |
| 692 | where "iszero z \<longleftrightarrow> z = 0" | |
| 693 | ||
| 694 | lemma iszero_0 [simp]: "iszero 0" | |
| 695 | by (simp add: iszero_def) | |
| 696 | ||
| 697 | lemma not_iszero_1 [simp]: "\<not> iszero 1" | |
| 698 | by (simp add: iszero_def) | |
| 699 | ||
| 700 | lemma not_iszero_Numeral1: "\<not> iszero Numeral1" | |
| 701 | by (simp add: numeral_One) | |
| 702 | ||
| 703 | lemma iszero_neg_numeral [simp]: | |
| 704 | "iszero (neg_numeral w) \<longleftrightarrow> iszero (numeral w)" | |
| 705 | unfolding iszero_def neg_numeral_def | |
| 706 | by (rule neg_equal_0_iff_equal) | |
| 707 | ||
| 708 | lemma eq_iff_iszero_diff: "x = y \<longleftrightarrow> iszero (x - y)" | |
| 709 | unfolding iszero_def by (rule eq_iff_diff_eq_0) | |
| 710 | ||
| 711 | text {* The @{text "eq_numeral_iff_iszero"} lemmas are not declared
 | |
| 712 | @{text "[simp]"} by default, because for rings of characteristic zero,
 | |
| 713 | better simp rules are possible. For a type like integers mod @{text
 | |
| 714 | "n"}, type-instantiated versions of these rules should be added to the | |
| 715 | simplifier, along with a type-specific rule for deciding propositions | |
| 716 | of the form @{text "iszero (numeral w)"}.
 | |
| 717 | ||
| 718 | bh: Maybe it would not be so bad to just declare these as simp | |
| 719 | rules anyway? I should test whether these rules take precedence over | |
| 720 | the @{text "ring_char_0"} rules in the simplifier.
 | |
| 721 | *} | |
| 722 | ||
| 723 | lemma eq_numeral_iff_iszero: | |
| 724 | "numeral x = numeral y \<longleftrightarrow> iszero (sub x y)" | |
| 725 | "numeral x = neg_numeral y \<longleftrightarrow> iszero (numeral (x + y))" | |
| 726 | "neg_numeral x = numeral y \<longleftrightarrow> iszero (numeral (x + y))" | |
| 727 | "neg_numeral x = neg_numeral y \<longleftrightarrow> iszero (sub y x)" | |
| 728 | "numeral x = 1 \<longleftrightarrow> iszero (sub x One)" | |
| 729 | "1 = numeral y \<longleftrightarrow> iszero (sub One y)" | |
| 730 | "neg_numeral x = 1 \<longleftrightarrow> iszero (numeral (x + One))" | |
| 731 | "1 = neg_numeral y \<longleftrightarrow> iszero (numeral (One + y))" | |
| 732 | "numeral x = 0 \<longleftrightarrow> iszero (numeral x)" | |
| 733 | "0 = numeral y \<longleftrightarrow> iszero (numeral y)" | |
| 734 | "neg_numeral x = 0 \<longleftrightarrow> iszero (numeral x)" | |
| 735 | "0 = neg_numeral y \<longleftrightarrow> iszero (numeral y)" | |
| 736 | unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special | |
| 737 | by simp_all | |
| 738 | ||
| 739 | end | |
| 740 | ||
| 741 | subsubsection {*
 | |
| 742 |   Equality and negation: class @{text ring_char_0}
 | |
| 743 | *} | |
| 744 | ||
| 745 | class ring_char_0 = ring_1 + semiring_char_0 | |
| 746 | begin | |
| 747 | ||
| 748 | lemma not_iszero_numeral [simp]: "\<not> iszero (numeral w)" | |
| 749 | by (simp add: iszero_def) | |
| 750 | ||
| 751 | lemma neg_numeral_eq_iff: "neg_numeral m = neg_numeral n \<longleftrightarrow> m = n" | |
| 752 | by (simp only: neg_numeral_def neg_equal_iff_equal numeral_eq_iff) | |
| 753 | ||
| 754 | lemma numeral_neq_neg_numeral: "numeral m \<noteq> neg_numeral n" | |
| 755 | unfolding neg_numeral_def eq_neg_iff_add_eq_0 | |
| 756 | by (simp add: numeral_plus_numeral) | |
| 757 | ||
| 758 | lemma neg_numeral_neq_numeral: "neg_numeral m \<noteq> numeral n" | |
| 759 | by (rule numeral_neq_neg_numeral [symmetric]) | |
| 760 | ||
| 761 | lemma zero_neq_neg_numeral: "0 \<noteq> neg_numeral n" | |
| 762 | unfolding neg_numeral_def neg_0_equal_iff_equal by simp | |
| 763 | ||
| 764 | lemma neg_numeral_neq_zero: "neg_numeral n \<noteq> 0" | |
| 765 | unfolding neg_numeral_def neg_equal_0_iff_equal by simp | |
| 766 | ||
| 767 | lemma one_neq_neg_numeral: "1 \<noteq> neg_numeral n" | |
| 768 | using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) | |
| 769 | ||
| 770 | lemma neg_numeral_neq_one: "neg_numeral n \<noteq> 1" | |
| 771 | using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) | |
| 772 | ||
| 773 | lemmas eq_neg_numeral_simps [simp] = | |
| 774 | neg_numeral_eq_iff | |
| 775 | numeral_neq_neg_numeral neg_numeral_neq_numeral | |
| 776 | one_neq_neg_numeral neg_numeral_neq_one | |
| 777 | zero_neq_neg_numeral neg_numeral_neq_zero | |
| 778 | ||
| 779 | end | |
| 780 | ||
| 781 | subsubsection {*
 | |
| 782 |   Structures with negation and order: class @{text linordered_idom}
 | |
| 783 | *} | |
| 784 | ||
| 785 | context linordered_idom | |
| 786 | begin | |
| 787 | ||
| 788 | subclass ring_char_0 .. | |
| 789 | ||
| 790 | lemma neg_numeral_le_iff: "neg_numeral m \<le> neg_numeral n \<longleftrightarrow> n \<le> m" | |
| 791 | by (simp only: neg_numeral_def neg_le_iff_le numeral_le_iff) | |
| 792 | ||
| 793 | lemma neg_numeral_less_iff: "neg_numeral m < neg_numeral n \<longleftrightarrow> n < m" | |
| 794 | by (simp only: neg_numeral_def neg_less_iff_less numeral_less_iff) | |
| 795 | ||
| 796 | lemma neg_numeral_less_zero: "neg_numeral n < 0" | |
| 797 | by (simp only: neg_numeral_def neg_less_0_iff_less zero_less_numeral) | |
| 798 | ||
| 799 | lemma neg_numeral_le_zero: "neg_numeral n \<le> 0" | |
| 800 | by (simp only: neg_numeral_def neg_le_0_iff_le zero_le_numeral) | |
| 801 | ||
| 802 | lemma not_zero_less_neg_numeral: "\<not> 0 < neg_numeral n" | |
| 803 | by (simp only: not_less neg_numeral_le_zero) | |
| 804 | ||
| 805 | lemma not_zero_le_neg_numeral: "\<not> 0 \<le> neg_numeral n" | |
| 806 | by (simp only: not_le neg_numeral_less_zero) | |
| 807 | ||
| 808 | lemma neg_numeral_less_numeral: "neg_numeral m < numeral n" | |
| 809 | using neg_numeral_less_zero zero_less_numeral by (rule less_trans) | |
| 810 | ||
| 811 | lemma neg_numeral_le_numeral: "neg_numeral m \<le> numeral n" | |
| 812 | by (simp only: less_imp_le neg_numeral_less_numeral) | |
| 813 | ||
| 814 | lemma not_numeral_less_neg_numeral: "\<not> numeral m < neg_numeral n" | |
| 815 | by (simp only: not_less neg_numeral_le_numeral) | |
| 816 | ||
| 817 | lemma not_numeral_le_neg_numeral: "\<not> numeral m \<le> neg_numeral n" | |
| 818 | by (simp only: not_le neg_numeral_less_numeral) | |
| 819 | ||
| 820 | lemma neg_numeral_less_one: "neg_numeral m < 1" | |
| 821 | by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) | |
| 822 | ||
| 823 | lemma neg_numeral_le_one: "neg_numeral m \<le> 1" | |
| 824 | by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) | |
| 825 | ||
| 826 | lemma not_one_less_neg_numeral: "\<not> 1 < neg_numeral m" | |
| 827 | by (simp only: not_less neg_numeral_le_one) | |
| 828 | ||
| 829 | lemma not_one_le_neg_numeral: "\<not> 1 \<le> neg_numeral m" | |
| 830 | by (simp only: not_le neg_numeral_less_one) | |
| 831 | ||
| 832 | lemma sub_non_negative: | |
| 833 | "sub n m \<ge> 0 \<longleftrightarrow> n \<ge> m" | |
| 834 | by (simp only: sub_def le_diff_eq) simp | |
| 835 | ||
| 836 | lemma sub_positive: | |
| 837 | "sub n m > 0 \<longleftrightarrow> n > m" | |
| 838 | by (simp only: sub_def less_diff_eq) simp | |
| 839 | ||
| 840 | lemma sub_non_positive: | |
| 841 | "sub n m \<le> 0 \<longleftrightarrow> n \<le> m" | |
| 842 | by (simp only: sub_def diff_le_eq) simp | |
| 843 | ||
| 844 | lemma sub_negative: | |
| 845 | "sub n m < 0 \<longleftrightarrow> n < m" | |
| 846 | by (simp only: sub_def diff_less_eq) simp | |
| 847 | ||
| 848 | lemmas le_neg_numeral_simps [simp] = | |
| 849 | neg_numeral_le_iff | |
| 850 | neg_numeral_le_numeral not_numeral_le_neg_numeral | |
| 851 | neg_numeral_le_zero not_zero_le_neg_numeral | |
| 852 | neg_numeral_le_one not_one_le_neg_numeral | |
| 853 | ||
| 854 | lemmas less_neg_numeral_simps [simp] = | |
| 855 | neg_numeral_less_iff | |
| 856 | neg_numeral_less_numeral not_numeral_less_neg_numeral | |
| 857 | neg_numeral_less_zero not_zero_less_neg_numeral | |
| 858 | neg_numeral_less_one not_one_less_neg_numeral | |
| 859 | ||
| 860 | lemma abs_numeral [simp]: "abs (numeral n) = numeral n" | |
| 861 | by simp | |
| 862 | ||
| 863 | lemma abs_neg_numeral [simp]: "abs (neg_numeral n) = numeral n" | |
| 864 | by (simp only: neg_numeral_def abs_minus_cancel abs_numeral) | |
| 865 | ||
| 866 | end | |
| 867 | ||
| 868 | subsubsection {*
 | |
| 869 | Natural numbers | |
| 870 | *} | |
| 871 | ||
| 47299 | 872 | lemma Suc_1 [simp]: "Suc 1 = 2" | 
| 873 | unfolding Suc_eq_plus1 by (rule one_add_one) | |
| 874 | ||
| 47108 | 875 | lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" | 
| 47299 | 876 | unfolding Suc_eq_plus1 by (rule numeral_plus_one) | 
| 47108 | 877 | |
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 878 | definition pred_numeral :: "num \<Rightarrow> nat" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 879 | where [code del]: "pred_numeral k = numeral k - 1" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 880 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 881 | lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 882 | unfolding pred_numeral_def by simp | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 883 | |
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 884 | lemma eval_nat_numeral: | 
| 47108 | 885 | "numeral One = Suc 0" | 
| 886 | "numeral (Bit0 n) = Suc (numeral (BitM n))" | |
| 887 | "numeral (Bit1 n) = Suc (numeral (Bit0 n))" | |
| 888 | by (simp_all add: numeral.simps BitM_plus_one) | |
| 889 | ||
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 890 | lemma pred_numeral_simps [simp]: | 
| 47300 | 891 | "pred_numeral One = 0" | 
| 892 | "pred_numeral (Bit0 k) = numeral (BitM k)" | |
| 893 | "pred_numeral (Bit1 k) = numeral (Bit0 k)" | |
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 894 | unfolding pred_numeral_def eval_nat_numeral | 
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 895 | by (simp_all only: diff_Suc_Suc diff_0) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 896 | |
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 897 | lemma numeral_2_eq_2: "2 = Suc (Suc 0)" | 
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 898 | by (simp add: eval_nat_numeral) | 
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 899 | |
| 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 900 | lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))" | 
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 901 | by (simp add: eval_nat_numeral) | 
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 902 | |
| 47207 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 903 | lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 904 | by (simp only: numeral_One One_nat_def) | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 905 | |
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 906 | lemma Suc_nat_number_of_add: | 
| 47300 | 907 | "Suc (numeral v + n) = numeral (v + One) + n" | 
| 47207 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 908 | by simp | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 909 | |
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 910 | (*Maps #n to n for n = 1, 2*) | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 911 | lemmas numerals = numeral_One [where 'a=nat] numeral_2_eq_2 | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 912 | |
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 913 | text {* Comparisons involving @{term Suc}. *}
 | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 914 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 915 | lemma eq_numeral_Suc [simp]: "numeral k = Suc n \<longleftrightarrow> pred_numeral k = n" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 916 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 917 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 918 | lemma Suc_eq_numeral [simp]: "Suc n = numeral k \<longleftrightarrow> n = pred_numeral k" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 919 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 920 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 921 | lemma less_numeral_Suc [simp]: "numeral k < Suc n \<longleftrightarrow> pred_numeral k < n" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 922 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 923 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 924 | lemma less_Suc_numeral [simp]: "Suc n < numeral k \<longleftrightarrow> n < pred_numeral k" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 925 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 926 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 927 | lemma le_numeral_Suc [simp]: "numeral k \<le> Suc n \<longleftrightarrow> pred_numeral k \<le> n" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 928 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 929 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 930 | lemma le_Suc_numeral [simp]: "Suc n \<le> numeral k \<longleftrightarrow> n \<le> pred_numeral k" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 931 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 932 | |
| 47218 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 933 | lemma diff_Suc_numeral [simp]: "Suc n - numeral k = n - pred_numeral k" | 
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 934 | by (simp add: numeral_eq_Suc) | 
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 935 | |
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 936 | lemma diff_numeral_Suc [simp]: "numeral k - Suc n = pred_numeral k - n" | 
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 937 | by (simp add: numeral_eq_Suc) | 
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 938 | |
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 939 | lemma max_Suc_numeral [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 940 | "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 941 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 942 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 943 | lemma max_numeral_Suc [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 944 | "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 945 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 946 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 947 | lemma min_Suc_numeral [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 948 | "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 949 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 950 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 951 | lemma min_numeral_Suc [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 952 | "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 953 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 954 | |
| 47216 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 955 | text {* For @{term nat_case} and @{term nat_rec}. *}
 | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 956 | |
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 957 | lemma nat_case_numeral [simp]: | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 958 | "nat_case a f (numeral v) = (let pv = pred_numeral v in f pv)" | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 959 | by (simp add: numeral_eq_Suc) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 960 | |
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 961 | lemma nat_case_add_eq_if [simp]: | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 962 | "nat_case a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))" | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 963 | by (simp add: numeral_eq_Suc) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 964 | |
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 965 | lemma nat_rec_numeral [simp]: | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 966 | "nat_rec a f (numeral v) = | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 967 | (let pv = pred_numeral v in f pv (nat_rec a f pv))" | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 968 | by (simp add: numeral_eq_Suc Let_def) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 969 | |
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 970 | lemma nat_rec_add_eq_if [simp]: | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 971 | "nat_rec a f (numeral v + n) = | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 972 | (let pv = pred_numeral v in f (pv + n) (nat_rec a f (pv + n)))" | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 973 | by (simp add: numeral_eq_Suc Let_def) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 974 | |
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 975 | text {* Case analysis on @{term "n < 2"} *}
 | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 976 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 977 | lemma less_2_cases: "n < 2 \<Longrightarrow> n = 0 \<or> n = Suc 0" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 978 | by (auto simp add: numeral_2_eq_2) | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 979 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 980 | text {* Removal of Small Numerals: 0, 1 and (in additive positions) 2 *}
 | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 981 | text {* bh: Are these rules really a good idea? *}
 | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 982 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 983 | lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 984 | by simp | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 985 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 986 | lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 987 | by simp | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 988 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 989 | text {* Can be used to eliminate long strings of Sucs, but not by default. *}
 | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 990 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 991 | lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 992 | by simp | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 993 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 994 | lemmas nat_1_add_1 = one_add_one [where 'a=nat] (* legacy *) | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 995 | |
| 47108 | 996 | |
| 997 | subsection {* Numeral equations as default simplification rules *}
 | |
| 998 | ||
| 999 | declare (in numeral) numeral_One [simp] | |
| 1000 | declare (in numeral) numeral_plus_numeral [simp] | |
| 1001 | declare (in numeral) add_numeral_special [simp] | |
| 1002 | declare (in neg_numeral) add_neg_numeral_simps [simp] | |
| 1003 | declare (in neg_numeral) add_neg_numeral_special [simp] | |
| 1004 | declare (in neg_numeral) diff_numeral_simps [simp] | |
| 1005 | declare (in neg_numeral) diff_numeral_special [simp] | |
| 1006 | declare (in semiring_numeral) numeral_times_numeral [simp] | |
| 1007 | declare (in ring_1) mult_neg_numeral_simps [simp] | |
| 1008 | ||
| 1009 | subsection {* Setting up simprocs *}
 | |
| 1010 | ||
| 1011 | lemma mult_numeral_1: "Numeral1 * a = (a::'a::semiring_numeral)" | |
| 1012 | by simp | |
| 1013 | ||
| 1014 | lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::semiring_numeral)" | |
| 1015 | by simp | |
| 1016 | ||
| 1017 | lemma divide_numeral_1: "a / Numeral1 = (a::'a::field)" | |
| 1018 | by simp | |
| 1019 | ||
| 1020 | lemma inverse_numeral_1: | |
| 1021 | "inverse Numeral1 = (Numeral1::'a::division_ring)" | |
| 1022 | by simp | |
| 1023 | ||
| 47211 | 1024 | text{*Theorem lists for the cancellation simprocs. The use of a binary
 | 
| 47108 | 1025 | numeral for 1 reduces the number of special cases.*} | 
| 1026 | ||
| 1027 | lemmas mult_1s = | |
| 1028 | mult_numeral_1 mult_numeral_1_right | |
| 1029 | mult_minus1 mult_minus1_right | |
| 1030 | ||
| 47226 | 1031 | setup {*
 | 
| 1032 | Reorient_Proc.add | |
| 1033 |     (fn Const (@{const_name numeral}, _) $ _ => true
 | |
| 1034 |     | Const (@{const_name neg_numeral}, _) $ _ => true
 | |
| 1035 | | _ => false) | |
| 1036 | *} | |
| 1037 | ||
| 1038 | simproc_setup reorient_numeral | |
| 1039 |   ("numeral w = x" | "neg_numeral w = y") = Reorient_Proc.proc
 | |
| 1040 | ||
| 47108 | 1041 | |
| 1042 | subsubsection {* Simplification of arithmetic operations on integer constants. *}
 | |
| 1043 | ||
| 1044 | lemmas arith_special = (* already declared simp above *) | |
| 1045 | add_numeral_special add_neg_numeral_special | |
| 1046 | diff_numeral_special minus_one | |
| 1047 | ||
| 1048 | (* rules already in simpset *) | |
| 1049 | lemmas arith_extra_simps = | |
| 1050 | numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right | |
| 1051 | minus_numeral minus_neg_numeral minus_zero minus_one | |
| 1052 | diff_numeral_simps diff_0 diff_0_right | |
| 1053 | numeral_times_numeral mult_neg_numeral_simps | |
| 1054 | mult_zero_left mult_zero_right | |
| 1055 | abs_numeral abs_neg_numeral | |
| 1056 | ||
| 1057 | text {*
 | |
| 1058 | For making a minimal simpset, one must include these default simprules. | |
| 1059 |   Also include @{text simp_thms}.
 | |
| 1060 | *} | |
| 1061 | ||
| 1062 | lemmas arith_simps = | |
| 1063 | add_num_simps mult_num_simps sub_num_simps | |
| 1064 | BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps | |
| 1065 | abs_zero abs_one arith_extra_simps | |
| 1066 | ||
| 1067 | text {* Simplification of relational operations *}
 | |
| 1068 | ||
| 1069 | lemmas eq_numeral_extra = | |
| 1070 | zero_neq_one one_neq_zero | |
| 1071 | ||
| 1072 | lemmas rel_simps = | |
| 1073 | le_num_simps less_num_simps eq_num_simps | |
| 1074 | le_numeral_simps le_neg_numeral_simps le_numeral_extra | |
| 1075 | less_numeral_simps less_neg_numeral_simps less_numeral_extra | |
| 1076 | eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra | |
| 1077 | ||
| 1078 | ||
| 1079 | subsubsection {* Simplification of arithmetic when nested to the right. *}
 | |
| 1080 | ||
| 1081 | lemma add_numeral_left [simp]: | |
| 1082 | "numeral v + (numeral w + z) = (numeral(v + w) + z)" | |
| 1083 | by (simp_all add: add_assoc [symmetric]) | |
| 1084 | ||
| 1085 | lemma add_neg_numeral_left [simp]: | |
| 1086 | "numeral v + (neg_numeral w + y) = (sub v w + y)" | |
| 1087 | "neg_numeral v + (numeral w + y) = (sub w v + y)" | |
| 1088 | "neg_numeral v + (neg_numeral w + y) = (neg_numeral(v + w) + y)" | |
| 1089 | by (simp_all add: add_assoc [symmetric]) | |
| 1090 | ||
| 1091 | lemma mult_numeral_left [simp]: | |
| 1092 | "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" | |
| 1093 | "neg_numeral v * (numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)" | |
| 1094 | "numeral v * (neg_numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)" | |
| 1095 | "neg_numeral v * (neg_numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)" | |
| 1096 | by (simp_all add: mult_assoc [symmetric]) | |
| 1097 | ||
| 1098 | hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec | |
| 1099 | ||
| 1100 | subsection {* code module namespace *}
 | |
| 1101 | ||
| 1102 | code_modulename SML | |
| 47126 | 1103 | Num Arith | 
| 47108 | 1104 | |
| 1105 | code_modulename OCaml | |
| 47126 | 1106 | Num Arith | 
| 47108 | 1107 | |
| 1108 | code_modulename Haskell | |
| 47126 | 1109 | Num Arith | 
| 47108 | 1110 | |
| 1111 | end | |
| 50817 | 1112 |