| author | wenzelm | 
| Mon, 11 Jul 2016 10:43:27 +0200 | |
| changeset 63433 | aa03b0487bf5 | 
| parent 62597 | b3f2b8c906a6 | 
| child 63654 | f90e3926e627 | 
| permissions | -rw-r--r-- | 
| 47108 | 1 | (* Title: HOL/Num.thy | 
| 2 | Author: Florian Haftmann | |
| 3 | Author: Brian Huffman | |
| 4 | *) | |
| 5 | ||
| 60758 | 6 | section \<open>Binary Numerals\<close> | 
| 47108 | 7 | |
| 8 | theory Num | |
| 58154 | 9 | imports BNF_Least_Fixpoint | 
| 47108 | 10 | begin | 
| 11 | ||
| 61799 | 12 | subsection \<open>The \<open>num\<close> type\<close> | 
| 47108 | 13 | |
| 58310 | 14 | datatype num = One | Bit0 num | Bit1 num | 
| 47108 | 15 | |
| 60758 | 16 | text \<open>Increment function for type @{typ num}\<close>
 | 
| 47108 | 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 | ||
| 60758 | 23 | text \<open>Converting between type @{typ num} and type @{typ nat}\<close>
 | 
| 47108 | 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 | ||
| 60758 | 47 | text \<open> | 
| 47108 | 48 |   Type @{typ num} is isomorphic to the strictly positive
 | 
| 49 | natural numbers. | |
| 60758 | 50 | \<close> | 
| 47108 | 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 | ||
| 60758 | 84 | text \<open> | 
| 47108 | 85 |   From now on, there are two possible models for @{typ num}:
 | 
| 61799 | 86 | as positive naturals (rule \<open>num_induct\<close>) | 
| 87 | and as digit representation (rules \<open>num.induct\<close>, \<open>num.cases\<close>). | |
| 60758 | 88 | \<close> | 
| 47108 | 89 | |
| 90 | ||
| 60758 | 91 | subsection \<open>Numeral operations\<close> | 
| 47108 | 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 | |
| 61169 | 109 | by standard (auto simp add: less_num_def less_eq_num_def num_eq_iff) | 
| 47108 | 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 | ||
| 61630 | 177 | lemma le_num_One_iff: "x \<le> num.One \<longleftrightarrow> x = num.One" | 
| 178 | by (simp add: antisym_conv) | |
| 179 | ||
| 61799 | 180 | text \<open>Rules using \<open>One\<close> and \<open>inc\<close> as constructors\<close> | 
| 47108 | 181 | |
| 182 | lemma add_One: "x + One = inc x" | |
| 183 | by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) | |
| 184 | ||
| 185 | lemma add_One_commute: "One + n = n + One" | |
| 186 | by (induct n) simp_all | |
| 187 | ||
| 188 | lemma add_inc: "x + inc y = inc (x + y)" | |
| 189 | by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) | |
| 190 | ||
| 191 | lemma mult_inc: "x * inc y = x * y + x" | |
| 192 | by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) | |
| 193 | ||
| 60758 | 194 | text \<open>The @{const num_of_nat} conversion\<close>
 | 
| 47108 | 195 | |
| 196 | lemma num_of_nat_One: | |
| 47300 | 197 | "n \<le> 1 \<Longrightarrow> num_of_nat n = One" | 
| 47108 | 198 | by (cases n) simp_all | 
| 199 | ||
| 200 | lemma num_of_nat_plus_distrib: | |
| 201 | "0 < m \<Longrightarrow> 0 < n \<Longrightarrow> num_of_nat (m + n) = num_of_nat m + num_of_nat n" | |
| 202 | by (induct n) (auto simp add: add_One add_One_commute add_inc) | |
| 203 | ||
| 60758 | 204 | text \<open>A double-and-decrement function\<close> | 
| 47108 | 205 | |
| 206 | primrec BitM :: "num \<Rightarrow> num" where | |
| 207 | "BitM One = One" | | |
| 208 | "BitM (Bit0 n) = Bit1 (BitM n)" | | |
| 209 | "BitM (Bit1 n) = Bit1 (Bit0 n)" | |
| 210 | ||
| 211 | lemma BitM_plus_one: "BitM n + One = Bit0 n" | |
| 212 | by (induct n) simp_all | |
| 213 | ||
| 214 | lemma one_plus_BitM: "One + BitM n = Bit0 n" | |
| 215 | unfolding add_One_commute BitM_plus_one .. | |
| 216 | ||
| 60758 | 217 | text \<open>Squaring and exponentiation\<close> | 
| 47108 | 218 | |
| 219 | primrec sqr :: "num \<Rightarrow> num" where | |
| 220 | "sqr One = One" | | |
| 221 | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | | |
| 222 | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" | |
| 223 | ||
| 224 | primrec pow :: "num \<Rightarrow> num \<Rightarrow> num" where | |
| 225 | "pow x One = x" | | |
| 226 | "pow x (Bit0 y) = sqr (pow x y)" | | |
| 47191 | 227 | "pow x (Bit1 y) = sqr (pow x y) * x" | 
| 47108 | 228 | |
| 229 | lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" | |
| 230 | by (induct x, simp_all add: algebra_simps nat_of_num_add) | |
| 231 | ||
| 232 | lemma sqr_conv_mult: "sqr x = x * x" | |
| 233 | by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) | |
| 234 | ||
| 235 | ||
| 60758 | 236 | subsection \<open>Binary numerals\<close> | 
| 47108 | 237 | |
| 60758 | 238 | text \<open> | 
| 47211 | 239 | We embed binary representations into a generic algebraic | 
| 61799 | 240 | structure using \<open>numeral\<close>. | 
| 60758 | 241 | \<close> | 
| 47108 | 242 | |
| 243 | class numeral = one + semigroup_add | |
| 244 | begin | |
| 245 | ||
| 246 | primrec numeral :: "num \<Rightarrow> 'a" where | |
| 247 | numeral_One: "numeral One = 1" | | |
| 248 | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | | |
| 249 | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" | |
| 250 | ||
| 50817 | 251 | lemma numeral_code [code]: | 
| 252 | "numeral One = 1" | |
| 253 | "numeral (Bit0 n) = (let m = numeral n in m + m)" | |
| 254 | "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" | |
| 255 | by (simp_all add: Let_def) | |
| 256 | ||
| 47108 | 257 | lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" | 
| 258 | apply (induct x) | |
| 259 | apply simp | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 260 | apply (simp add: add.assoc [symmetric], simp add: add.assoc) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 261 | apply (simp add: add.assoc [symmetric], simp add: add.assoc) | 
| 47108 | 262 | done | 
| 263 | ||
| 264 | lemma numeral_inc: "numeral (inc x) = numeral x + 1" | |
| 265 | proof (induct x) | |
| 266 | case (Bit1 x) | |
| 267 | have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" | |
| 268 | by (simp only: one_plus_numeral_commute) | |
| 269 | with Bit1 show ?case | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 270 | by (simp add: add.assoc) | 
| 47108 | 271 | qed simp_all | 
| 272 | ||
| 273 | declare numeral.simps [simp del] | |
| 274 | ||
| 275 | abbreviation "Numeral1 \<equiv> numeral One" | |
| 276 | ||
| 277 | declare numeral_One [code_post] | |
| 278 | ||
| 279 | end | |
| 280 | ||
| 60758 | 281 | text \<open>Numeral syntax.\<close> | 
| 47108 | 282 | |
| 283 | syntax | |
| 284 |   "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
 | |
| 285 | ||
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 286 | ML_file "Tools/numeral.ML" | 
| 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 287 | |
| 60758 | 288 | parse_translation \<open> | 
| 52143 | 289 | let | 
| 290 |     fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
 | |
| 291 | c $ numeral_tr [t] $ u | |
| 292 | | numeral_tr [Const (num, _)] = | |
| 58421 | 293 | (Numeral.mk_number_syntax o #value o Lexicon.read_num) num | 
| 52143 | 294 |       | numeral_tr ts = raise TERM ("numeral_tr", ts);
 | 
| 55974 
c835a9379026
more official const syntax: avoid educated guessing by Syntax_Phases.decode_term;
 wenzelm parents: 
55534diff
changeset | 295 |   in [(@{syntax_const "_Numeral"}, K numeral_tr)] end
 | 
| 60758 | 296 | \<close> | 
| 47108 | 297 | |
| 60758 | 298 | typed_print_translation \<open> | 
| 52143 | 299 | let | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 300 | fun num_tr' ctxt T [n] = | 
| 52143 | 301 | let | 
| 62597 | 302 | val k = Numeral.dest_num_syntax n; | 
| 52187 | 303 | val t' = | 
| 304 |           Syntax.const @{syntax_const "_Numeral"} $
 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 305 | Syntax.free (string_of_int k); | 
| 52143 | 306 | in | 
| 307 | (case T of | |
| 308 |           Type (@{type_name fun}, [_, T']) =>
 | |
| 52210 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52187diff
changeset | 309 | if Printer.type_emphasis ctxt T' then | 
| 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52187diff
changeset | 310 |               Syntax.const @{syntax_const "_constrain"} $ t' $
 | 
| 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52187diff
changeset | 311 | Syntax_Phases.term_of_typ ctxt T' | 
| 
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
 wenzelm parents: 
52187diff
changeset | 312 | else t' | 
| 52187 | 313 | | _ => if T = dummyT then t' else raise Match) | 
| 52143 | 314 | end; | 
| 315 | in | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 316 |    [(@{const_syntax numeral}, num_tr')]
 | 
| 52143 | 317 | end | 
| 60758 | 318 | \<close> | 
| 47108 | 319 | |
| 47228 | 320 | |
| 60758 | 321 | subsection \<open>Class-specific numeral rules\<close> | 
| 47108 | 322 | |
| 60758 | 323 | text \<open> | 
| 47108 | 324 |   @{const numeral} is a morphism.
 | 
| 60758 | 325 | \<close> | 
| 47108 | 326 | |
| 61799 | 327 | subsubsection \<open>Structures with addition: class \<open>numeral\<close>\<close> | 
| 47108 | 328 | |
| 329 | context numeral | |
| 330 | begin | |
| 331 | ||
| 332 | lemma numeral_add: "numeral (m + n) = numeral m + numeral n" | |
| 333 | by (induct n rule: num_induct) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 334 | (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) | 
| 47108 | 335 | |
| 336 | lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" | |
| 337 | by (rule numeral_add [symmetric]) | |
| 338 | ||
| 339 | lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" | |
| 340 | using numeral_add [of n One] by (simp add: numeral_One) | |
| 341 | ||
| 342 | lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" | |
| 343 | using numeral_add [of One n] by (simp add: numeral_One) | |
| 344 | ||
| 345 | lemma one_add_one: "1 + 1 = 2" | |
| 346 | using numeral_add [of One One] by (simp add: numeral_One) | |
| 347 | ||
| 348 | lemmas add_numeral_special = | |
| 349 | numeral_plus_one one_plus_numeral one_add_one | |
| 350 | ||
| 351 | end | |
| 352 | ||
| 60758 | 353 | subsubsection \<open> | 
| 61799 | 354 | Structures with negation: class \<open>neg_numeral\<close> | 
| 60758 | 355 | \<close> | 
| 47108 | 356 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 357 | class neg_numeral = numeral + group_add | 
| 47108 | 358 | begin | 
| 359 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 360 | lemma uminus_numeral_One: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 361 | "- Numeral1 = - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 362 | by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 363 | |
| 60758 | 364 | text \<open>Numerals form an abelian subgroup.\<close> | 
| 47108 | 365 | |
| 366 | inductive is_num :: "'a \<Rightarrow> bool" where | |
| 367 | "is_num 1" | | |
| 368 | "is_num x \<Longrightarrow> is_num (- x)" | | |
| 369 | "\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> is_num (x + y)" | |
| 370 | ||
| 371 | lemma is_num_numeral: "is_num (numeral k)" | |
| 372 | by (induct k, simp_all add: numeral.simps is_num.intros) | |
| 373 | ||
| 374 | lemma is_num_add_commute: | |
| 375 | "\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> x + y = y + x" | |
| 376 | apply (induct x rule: is_num.induct) | |
| 377 | apply (induct y rule: is_num.induct) | |
| 378 | apply simp | |
| 379 | apply (rule_tac a=x in add_left_imp_eq) | |
| 380 | apply (rule_tac a=x in add_right_imp_eq) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 381 | apply (simp add: add.assoc) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 382 | apply (simp add: add.assoc [symmetric], simp add: add.assoc) | 
| 47108 | 383 | apply (rule_tac a=x in add_left_imp_eq) | 
| 384 | apply (rule_tac a=x in add_right_imp_eq) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 385 | apply (simp add: add.assoc) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 386 | apply (simp add: add.assoc, simp add: add.assoc [symmetric]) | 
| 47108 | 387 | done | 
| 388 | ||
| 389 | lemma is_num_add_left_commute: | |
| 390 | "\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> x + (y + z) = y + (x + z)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 391 | by (simp only: add.assoc [symmetric] is_num_add_commute) | 
| 47108 | 392 | |
| 393 | lemmas is_num_normalize = | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 394 | add.assoc is_num_add_commute is_num_add_left_commute | 
| 47108 | 395 | is_num.intros is_num_numeral | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53064diff
changeset | 396 | minus_add | 
| 47108 | 397 | |
| 398 | definition dbl :: "'a \<Rightarrow> 'a" where "dbl x = x + x" | |
| 399 | definition dbl_inc :: "'a \<Rightarrow> 'a" where "dbl_inc x = x + x + 1" | |
| 400 | definition dbl_dec :: "'a \<Rightarrow> 'a" where "dbl_dec x = x + x - 1" | |
| 401 | ||
| 402 | definition sub :: "num \<Rightarrow> num \<Rightarrow> 'a" where | |
| 403 | "sub k l = numeral k - numeral l" | |
| 404 | ||
| 405 | lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" | |
| 406 | by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) | |
| 407 | ||
| 408 | lemma dbl_simps [simp]: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 409 | "dbl (- numeral k) = - dbl (numeral k)" | 
| 47108 | 410 | "dbl 0 = 0" | 
| 411 | "dbl 1 = 2" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 412 | "dbl (- 1) = - 2" | 
| 47108 | 413 | "dbl (numeral k) = numeral (Bit0 k)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 414 | by (simp_all add: dbl_def numeral.simps minus_add) | 
| 47108 | 415 | |
| 416 | lemma dbl_inc_simps [simp]: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 417 | "dbl_inc (- numeral k) = - dbl_dec (numeral k)" | 
| 47108 | 418 | "dbl_inc 0 = 1" | 
| 419 | "dbl_inc 1 = 3" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 420 | "dbl_inc (- 1) = - 1" | 
| 47108 | 421 | "dbl_inc (numeral k) = numeral (Bit1 k)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 422 | by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff) | 
| 47108 | 423 | |
| 424 | lemma dbl_dec_simps [simp]: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 425 | "dbl_dec (- numeral k) = - dbl_inc (numeral k)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 426 | "dbl_dec 0 = - 1" | 
| 47108 | 427 | "dbl_dec 1 = 1" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 428 | "dbl_dec (- 1) = - 3" | 
| 47108 | 429 | "dbl_dec (numeral k) = numeral (BitM k)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 430 | by (simp_all add: dbl_dec_def dbl_inc_def numeral.simps numeral_BitM is_num_normalize) | 
| 47108 | 431 | |
| 432 | lemma sub_num_simps [simp]: | |
| 433 | "sub One One = 0" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 434 | "sub One (Bit0 l) = - numeral (BitM l)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 435 | "sub One (Bit1 l) = - numeral (Bit0 l)" | 
| 47108 | 436 | "sub (Bit0 k) One = numeral (BitM k)" | 
| 437 | "sub (Bit1 k) One = numeral (Bit0 k)" | |
| 438 | "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" | |
| 439 | "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" | |
| 440 | "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" | |
| 441 | "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 442 | by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def numeral.simps | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53064diff
changeset | 443 | numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) | 
| 47108 | 444 | |
| 445 | lemma add_neg_numeral_simps: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 446 | "numeral m + - numeral n = sub m n" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 447 | "- numeral m + numeral n = sub n m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 448 | "- numeral m + - numeral n = - (numeral m + numeral n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 449 | by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53064diff
changeset | 450 | del: add_uminus_conv_diff add: diff_conv_add_uminus) | 
| 47108 | 451 | |
| 452 | lemma add_neg_numeral_special: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 453 | "1 + - numeral m = sub One m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 454 | "- numeral m + 1 = sub One m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 455 | "numeral m + - 1 = sub m One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 456 | "- 1 + numeral n = sub n One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 457 | "- 1 + - numeral n = - numeral (inc n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 458 | "- numeral m + - 1 = - numeral (inc m)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 459 | "1 + - 1 = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 460 | "- 1 + 1 = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 461 | "- 1 + - 1 = - 2" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 462 | by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 463 | del: add_uminus_conv_diff add: diff_conv_add_uminus) | 
| 47108 | 464 | |
| 465 | lemma diff_numeral_simps: | |
| 466 | "numeral m - numeral n = sub m n" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 467 | "numeral m - - numeral n = numeral (m + n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 468 | "- numeral m - numeral n = - numeral (m + n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 469 | "- numeral m - - numeral n = sub n m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 470 | by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53064diff
changeset | 471 | del: add_uminus_conv_diff add: diff_conv_add_uminus) | 
| 47108 | 472 | |
| 473 | lemma diff_numeral_special: | |
| 474 | "1 - numeral n = sub One n" | |
| 475 | "numeral m - 1 = sub m One" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 476 | "1 - - numeral n = numeral (One + n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 477 | "- numeral m - 1 = - numeral (m + One)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 478 | "- 1 - numeral n = - numeral (inc n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 479 | "numeral m - - 1 = numeral (inc m)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 480 | "- 1 - - numeral n = sub n One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 481 | "- numeral m - - 1 = sub One m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 482 | "1 - 1 = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 483 | "- 1 - 1 = - 2" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 484 | "1 - - 1 = 2" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 485 | "- 1 - - 1 = 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 486 | by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 487 | del: add_uminus_conv_diff add: diff_conv_add_uminus) | 
| 47108 | 488 | |
| 489 | end | |
| 490 | ||
| 60758 | 491 | subsubsection \<open> | 
| 61799 | 492 | Structures with multiplication: class \<open>semiring_numeral\<close> | 
| 60758 | 493 | \<close> | 
| 47108 | 494 | |
| 495 | class semiring_numeral = semiring + monoid_mult | |
| 496 | begin | |
| 497 | ||
| 498 | subclass numeral .. | |
| 499 | ||
| 500 | lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" | |
| 501 | apply (induct n rule: num_induct) | |
| 502 | apply (simp add: numeral_One) | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49690diff
changeset | 503 | apply (simp add: mult_inc numeral_inc numeral_add distrib_left) | 
| 47108 | 504 | done | 
| 505 | ||
| 506 | lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" | |
| 507 | by (rule numeral_mult [symmetric]) | |
| 508 | ||
| 53064 | 509 | lemma mult_2: "2 * z = z + z" | 
| 510 | unfolding one_add_one [symmetric] distrib_right by simp | |
| 511 | ||
| 512 | lemma mult_2_right: "z * 2 = z + z" | |
| 513 | unfolding one_add_one [symmetric] distrib_left by simp | |
| 514 | ||
| 47108 | 515 | end | 
| 516 | ||
| 60758 | 517 | subsubsection \<open> | 
| 61799 | 518 | Structures with a zero: class \<open>semiring_1\<close> | 
| 60758 | 519 | \<close> | 
| 47108 | 520 | |
| 521 | context semiring_1 | |
| 522 | begin | |
| 523 | ||
| 524 | subclass semiring_numeral .. | |
| 525 | ||
| 526 | lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" | |
| 527 | by (induct n, | |
| 528 | simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) | |
| 529 | ||
| 530 | end | |
| 531 | ||
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 532 | lemma nat_of_num_numeral [code_abbrev]: | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 533 | "nat_of_num = numeral" | 
| 47108 | 534 | proof | 
| 535 | fix n | |
| 536 | have "numeral n = nat_of_num n" | |
| 537 | by (induct n) (simp_all add: numeral.simps) | |
| 538 | then show "nat_of_num n = numeral n" by simp | |
| 539 | qed | |
| 540 | ||
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 541 | lemma nat_of_num_code [code]: | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 542 | "nat_of_num One = 1" | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 543 | "nat_of_num (Bit0 n) = (let m = nat_of_num n in m + m)" | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 544 | "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))" | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 545 | by (simp_all add: Let_def) | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 546 | |
| 60758 | 547 | subsubsection \<open> | 
| 61799 | 548 | Equality: class \<open>semiring_char_0\<close> | 
| 60758 | 549 | \<close> | 
| 47108 | 550 | |
| 551 | context semiring_char_0 | |
| 552 | begin | |
| 553 | ||
| 554 | lemma numeral_eq_iff: "numeral m = numeral n \<longleftrightarrow> m = n" | |
| 555 | unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] | |
| 556 | of_nat_eq_iff num_eq_iff .. | |
| 557 | ||
| 558 | lemma numeral_eq_one_iff: "numeral n = 1 \<longleftrightarrow> n = One" | |
| 559 | by (rule numeral_eq_iff [of n One, unfolded numeral_One]) | |
| 560 | ||
| 561 | lemma one_eq_numeral_iff: "1 = numeral n \<longleftrightarrow> One = n" | |
| 562 | by (rule numeral_eq_iff [of One n, unfolded numeral_One]) | |
| 563 | ||
| 564 | lemma numeral_neq_zero: "numeral n \<noteq> 0" | |
| 565 | unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] | |
| 566 | by (simp add: nat_of_num_pos) | |
| 567 | ||
| 568 | lemma zero_neq_numeral: "0 \<noteq> numeral n" | |
| 569 | unfolding eq_commute [of 0] by (rule numeral_neq_zero) | |
| 570 | ||
| 571 | lemmas eq_numeral_simps [simp] = | |
| 572 | numeral_eq_iff | |
| 573 | numeral_eq_one_iff | |
| 574 | one_eq_numeral_iff | |
| 575 | numeral_neq_zero | |
| 576 | zero_neq_numeral | |
| 577 | ||
| 578 | end | |
| 579 | ||
| 60758 | 580 | subsubsection \<open> | 
| 61799 | 581 | Comparisons: class \<open>linordered_semidom\<close> | 
| 60758 | 582 | \<close> | 
| 47108 | 583 | |
| 60758 | 584 | text \<open>Could be perhaps more general than here.\<close> | 
| 47108 | 585 | |
| 586 | context linordered_semidom | |
| 587 | begin | |
| 588 | ||
| 589 | lemma numeral_le_iff: "numeral m \<le> numeral n \<longleftrightarrow> m \<le> n" | |
| 590 | proof - | |
| 591 | have "of_nat (numeral m) \<le> of_nat (numeral n) \<longleftrightarrow> m \<le> n" | |
| 592 | unfolding less_eq_num_def nat_of_num_numeral of_nat_le_iff .. | |
| 593 | then show ?thesis by simp | |
| 594 | qed | |
| 595 | ||
| 596 | lemma one_le_numeral: "1 \<le> numeral n" | |
| 597 | using numeral_le_iff [of One n] by (simp add: numeral_One) | |
| 598 | ||
| 599 | lemma numeral_le_one_iff: "numeral n \<le> 1 \<longleftrightarrow> n \<le> One" | |
| 600 | using numeral_le_iff [of n One] by (simp add: numeral_One) | |
| 601 | ||
| 602 | lemma numeral_less_iff: "numeral m < numeral n \<longleftrightarrow> m < n" | |
| 603 | proof - | |
| 604 | have "of_nat (numeral m) < of_nat (numeral n) \<longleftrightarrow> m < n" | |
| 605 | unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. | |
| 606 | then show ?thesis by simp | |
| 607 | qed | |
| 608 | ||
| 609 | lemma not_numeral_less_one: "\<not> numeral n < 1" | |
| 610 | using numeral_less_iff [of n One] by (simp add: numeral_One) | |
| 611 | ||
| 612 | lemma one_less_numeral_iff: "1 < numeral n \<longleftrightarrow> One < n" | |
| 613 | using numeral_less_iff [of One n] by (simp add: numeral_One) | |
| 614 | ||
| 615 | lemma zero_le_numeral: "0 \<le> numeral n" | |
| 616 | by (induct n) (simp_all add: numeral.simps) | |
| 617 | ||
| 618 | lemma zero_less_numeral: "0 < numeral n" | |
| 619 | by (induct n) (simp_all add: numeral.simps add_pos_pos) | |
| 620 | ||
| 621 | lemma not_numeral_le_zero: "\<not> numeral n \<le> 0" | |
| 622 | by (simp add: not_le zero_less_numeral) | |
| 623 | ||
| 624 | lemma not_numeral_less_zero: "\<not> numeral n < 0" | |
| 625 | by (simp add: not_less zero_le_numeral) | |
| 626 | ||
| 627 | lemmas le_numeral_extra = | |
| 628 | zero_le_one not_one_le_zero | |
| 629 | order_refl [of 0] order_refl [of 1] | |
| 630 | ||
| 631 | lemmas less_numeral_extra = | |
| 632 | zero_less_one not_one_less_zero | |
| 633 | less_irrefl [of 0] less_irrefl [of 1] | |
| 634 | ||
| 635 | lemmas le_numeral_simps [simp] = | |
| 636 | numeral_le_iff | |
| 637 | one_le_numeral | |
| 638 | numeral_le_one_iff | |
| 639 | zero_le_numeral | |
| 640 | not_numeral_le_zero | |
| 641 | ||
| 642 | lemmas less_numeral_simps [simp] = | |
| 643 | numeral_less_iff | |
| 644 | one_less_numeral_iff | |
| 645 | not_numeral_less_one | |
| 646 | zero_less_numeral | |
| 647 | not_numeral_less_zero | |
| 648 | ||
| 61630 | 649 | lemma min_0_1 [simp]: | 
| 650 | fixes min' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" defines "min' \<equiv> min" shows | |
| 651 | "min' 0 1 = 0" | |
| 652 | "min' 1 0 = 0" | |
| 653 | "min' 0 (numeral x) = 0" | |
| 654 | "min' (numeral x) 0 = 0" | |
| 655 | "min' 1 (numeral x) = 1" | |
| 656 | "min' (numeral x) 1 = 1" | |
| 657 | by(simp_all add: min'_def min_def le_num_One_iff) | |
| 658 | ||
| 659 | lemma max_0_1 [simp]: | |
| 660 | fixes max' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" defines "max' \<equiv> max" shows | |
| 661 | "max' 0 1 = 1" | |
| 662 | "max' 1 0 = 1" | |
| 663 | "max' 0 (numeral x) = numeral x" | |
| 664 | "max' (numeral x) 0 = numeral x" | |
| 665 | "max' 1 (numeral x) = numeral x" | |
| 666 | "max' (numeral x) 1 = numeral x" | |
| 667 | by(simp_all add: max'_def max_def le_num_One_iff) | |
| 668 | ||
| 47108 | 669 | end | 
| 670 | ||
| 60758 | 671 | subsubsection \<open> | 
| 61799 | 672 | Multiplication and negation: class \<open>ring_1\<close> | 
| 60758 | 673 | \<close> | 
| 47108 | 674 | |
| 675 | context ring_1 | |
| 676 | begin | |
| 677 | ||
| 678 | subclass neg_numeral .. | |
| 679 | ||
| 680 | lemma mult_neg_numeral_simps: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 681 | "- numeral m * - numeral n = numeral (m * n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 682 | "- numeral m * numeral n = - numeral (m * n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 683 | "numeral m * - numeral n = - numeral (m * n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 684 | unfolding mult_minus_left mult_minus_right | 
| 47108 | 685 | by (simp_all only: minus_minus numeral_mult) | 
| 686 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 687 | lemma mult_minus1 [simp]: "- 1 * z = - z" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 688 | unfolding numeral.simps mult_minus_left by simp | 
| 47108 | 689 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 690 | lemma mult_minus1_right [simp]: "z * - 1 = - z" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 691 | unfolding numeral.simps mult_minus_right by simp | 
| 47108 | 692 | |
| 693 | end | |
| 694 | ||
| 60758 | 695 | subsubsection \<open> | 
| 61799 | 696 | Equality using \<open>iszero\<close> for rings with non-zero characteristic | 
| 60758 | 697 | \<close> | 
| 47108 | 698 | |
| 699 | context ring_1 | |
| 700 | begin | |
| 701 | ||
| 702 | definition iszero :: "'a \<Rightarrow> bool" | |
| 703 | where "iszero z \<longleftrightarrow> z = 0" | |
| 704 | ||
| 705 | lemma iszero_0 [simp]: "iszero 0" | |
| 706 | by (simp add: iszero_def) | |
| 707 | ||
| 708 | lemma not_iszero_1 [simp]: "\<not> iszero 1" | |
| 709 | by (simp add: iszero_def) | |
| 710 | ||
| 711 | lemma not_iszero_Numeral1: "\<not> iszero Numeral1" | |
| 712 | by (simp add: numeral_One) | |
| 713 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 714 | lemma not_iszero_neg_1 [simp]: "\<not> iszero (- 1)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 715 | by (simp add: iszero_def) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 716 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 717 | lemma not_iszero_neg_Numeral1: "\<not> iszero (- Numeral1)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 718 | by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 719 | |
| 47108 | 720 | lemma iszero_neg_numeral [simp]: | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 721 | "iszero (- numeral w) \<longleftrightarrow> iszero (numeral w)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 722 | unfolding iszero_def | 
| 47108 | 723 | by (rule neg_equal_0_iff_equal) | 
| 724 | ||
| 725 | lemma eq_iff_iszero_diff: "x = y \<longleftrightarrow> iszero (x - y)" | |
| 726 | unfolding iszero_def by (rule eq_iff_diff_eq_0) | |
| 727 | ||
| 61799 | 728 | text \<open>The \<open>eq_numeral_iff_iszero\<close> lemmas are not declared | 
| 729 | \<open>[simp]\<close> by default, because for rings of characteristic zero, | |
| 730 | better simp rules are possible. For a type like integers mod \<open>n\<close>, type-instantiated versions of these rules should be added to the | |
| 47108 | 731 | simplifier, along with a type-specific rule for deciding propositions | 
| 61799 | 732 | of the form \<open>iszero (numeral w)\<close>. | 
| 47108 | 733 | |
| 734 | bh: Maybe it would not be so bad to just declare these as simp | |
| 735 | rules anyway? I should test whether these rules take precedence over | |
| 61799 | 736 | the \<open>ring_char_0\<close> rules in the simplifier. | 
| 60758 | 737 | \<close> | 
| 47108 | 738 | |
| 739 | lemma eq_numeral_iff_iszero: | |
| 740 | "numeral x = numeral y \<longleftrightarrow> iszero (sub x y)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 741 | "numeral x = - numeral y \<longleftrightarrow> iszero (numeral (x + y))" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 742 | "- numeral x = numeral y \<longleftrightarrow> iszero (numeral (x + y))" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 743 | "- numeral x = - numeral y \<longleftrightarrow> iszero (sub y x)" | 
| 47108 | 744 | "numeral x = 1 \<longleftrightarrow> iszero (sub x One)" | 
| 745 | "1 = numeral y \<longleftrightarrow> iszero (sub One y)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 746 | "- numeral x = 1 \<longleftrightarrow> iszero (numeral (x + One))" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 747 | "1 = - numeral y \<longleftrightarrow> iszero (numeral (One + y))" | 
| 47108 | 748 | "numeral x = 0 \<longleftrightarrow> iszero (numeral x)" | 
| 749 | "0 = numeral y \<longleftrightarrow> iszero (numeral y)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 750 | "- numeral x = 0 \<longleftrightarrow> iszero (numeral x)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 751 | "0 = - numeral y \<longleftrightarrow> iszero (numeral y)" | 
| 47108 | 752 | unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special | 
| 753 | by simp_all | |
| 754 | ||
| 755 | end | |
| 756 | ||
| 60758 | 757 | subsubsection \<open> | 
| 61799 | 758 | Equality and negation: class \<open>ring_char_0\<close> | 
| 60758 | 759 | \<close> | 
| 47108 | 760 | |
| 62481 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 haftmann parents: 
62348diff
changeset | 761 | context ring_char_0 | 
| 47108 | 762 | begin | 
| 763 | ||
| 764 | lemma not_iszero_numeral [simp]: "\<not> iszero (numeral w)" | |
| 765 | by (simp add: iszero_def) | |
| 766 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 767 | lemma neg_numeral_eq_iff: "- numeral m = - numeral n \<longleftrightarrow> m = n" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 768 | by simp | 
| 47108 | 769 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 770 | lemma numeral_neq_neg_numeral: "numeral m \<noteq> - numeral n" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 771 | unfolding eq_neg_iff_add_eq_0 | 
| 47108 | 772 | by (simp add: numeral_plus_numeral) | 
| 773 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 774 | lemma neg_numeral_neq_numeral: "- numeral m \<noteq> numeral n" | 
| 47108 | 775 | by (rule numeral_neq_neg_numeral [symmetric]) | 
| 776 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 777 | lemma zero_neq_neg_numeral: "0 \<noteq> - numeral n" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 778 | unfolding neg_0_equal_iff_equal by simp | 
| 47108 | 779 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 780 | lemma neg_numeral_neq_zero: "- numeral n \<noteq> 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 781 | unfolding neg_equal_0_iff_equal by simp | 
| 47108 | 782 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 783 | lemma one_neq_neg_numeral: "1 \<noteq> - numeral n" | 
| 47108 | 784 | using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) | 
| 785 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 786 | lemma neg_numeral_neq_one: "- numeral n \<noteq> 1" | 
| 47108 | 787 | using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) | 
| 788 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 789 | lemma neg_one_neq_numeral: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 790 | "- 1 \<noteq> numeral n" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 791 | using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 792 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 793 | lemma numeral_neq_neg_one: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 794 | "numeral n \<noteq> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 795 | using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 796 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 797 | lemma neg_one_eq_numeral_iff: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 798 | "- 1 = - numeral n \<longleftrightarrow> n = One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 799 | using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 800 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 801 | lemma numeral_eq_neg_one_iff: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 802 | "- numeral n = - 1 \<longleftrightarrow> n = One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 803 | using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 804 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 805 | lemma neg_one_neq_zero: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 806 | "- 1 \<noteq> 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 807 | by simp | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 808 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 809 | lemma zero_neq_neg_one: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 810 | "0 \<noteq> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 811 | by simp | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 812 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 813 | lemma neg_one_neq_one: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 814 | "- 1 \<noteq> 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 815 | using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 816 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 817 | lemma one_neq_neg_one: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 818 | "1 \<noteq> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 819 | using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 820 | |
| 47108 | 821 | lemmas eq_neg_numeral_simps [simp] = | 
| 822 | neg_numeral_eq_iff | |
| 823 | numeral_neq_neg_numeral neg_numeral_neq_numeral | |
| 824 | one_neq_neg_numeral neg_numeral_neq_one | |
| 825 | zero_neq_neg_numeral neg_numeral_neq_zero | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 826 | neg_one_neq_numeral numeral_neq_neg_one | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 827 | neg_one_eq_numeral_iff numeral_eq_neg_one_iff | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 828 | neg_one_neq_zero zero_neq_neg_one | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 829 | neg_one_neq_one one_neq_neg_one | 
| 47108 | 830 | |
| 831 | end | |
| 832 | ||
| 62348 | 833 | |
| 60758 | 834 | subsubsection \<open> | 
| 61799 | 835 | Structures with negation and order: class \<open>linordered_idom\<close> | 
| 60758 | 836 | \<close> | 
| 47108 | 837 | |
| 838 | context linordered_idom | |
| 839 | begin | |
| 840 | ||
| 841 | subclass ring_char_0 .. | |
| 842 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 843 | lemma neg_numeral_le_iff: "- numeral m \<le> - numeral n \<longleftrightarrow> n \<le> m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 844 | by (simp only: neg_le_iff_le numeral_le_iff) | 
| 47108 | 845 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 846 | lemma neg_numeral_less_iff: "- numeral m < - numeral n \<longleftrightarrow> n < m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 847 | by (simp only: neg_less_iff_less numeral_less_iff) | 
| 47108 | 848 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 849 | lemma neg_numeral_less_zero: "- numeral n < 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 850 | by (simp only: neg_less_0_iff_less zero_less_numeral) | 
| 47108 | 851 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 852 | lemma neg_numeral_le_zero: "- numeral n \<le> 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 853 | by (simp only: neg_le_0_iff_le zero_le_numeral) | 
| 47108 | 854 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 855 | lemma not_zero_less_neg_numeral: "\<not> 0 < - numeral n" | 
| 47108 | 856 | by (simp only: not_less neg_numeral_le_zero) | 
| 857 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 858 | lemma not_zero_le_neg_numeral: "\<not> 0 \<le> - numeral n" | 
| 47108 | 859 | by (simp only: not_le neg_numeral_less_zero) | 
| 860 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 861 | lemma neg_numeral_less_numeral: "- numeral m < numeral n" | 
| 47108 | 862 | using neg_numeral_less_zero zero_less_numeral by (rule less_trans) | 
| 863 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 864 | lemma neg_numeral_le_numeral: "- numeral m \<le> numeral n" | 
| 47108 | 865 | by (simp only: less_imp_le neg_numeral_less_numeral) | 
| 866 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 867 | lemma not_numeral_less_neg_numeral: "\<not> numeral m < - numeral n" | 
| 47108 | 868 | by (simp only: not_less neg_numeral_le_numeral) | 
| 869 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 870 | lemma not_numeral_le_neg_numeral: "\<not> numeral m \<le> - numeral n" | 
| 47108 | 871 | by (simp only: not_le neg_numeral_less_numeral) | 
| 872 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 873 | lemma neg_numeral_less_one: "- numeral m < 1" | 
| 47108 | 874 | by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) | 
| 875 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 876 | lemma neg_numeral_le_one: "- numeral m \<le> 1" | 
| 47108 | 877 | by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) | 
| 878 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 879 | lemma not_one_less_neg_numeral: "\<not> 1 < - numeral m" | 
| 47108 | 880 | by (simp only: not_less neg_numeral_le_one) | 
| 881 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 882 | lemma not_one_le_neg_numeral: "\<not> 1 \<le> - numeral m" | 
| 47108 | 883 | by (simp only: not_le neg_numeral_less_one) | 
| 884 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 885 | lemma not_numeral_less_neg_one: "\<not> numeral m < - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 886 | using not_numeral_less_neg_numeral [of m One] by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 887 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 888 | lemma not_numeral_le_neg_one: "\<not> numeral m \<le> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 889 | using not_numeral_le_neg_numeral [of m One] by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 890 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 891 | lemma neg_one_less_numeral: "- 1 < numeral m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 892 | using neg_numeral_less_numeral [of One m] by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 893 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 894 | lemma neg_one_le_numeral: "- 1 \<le> numeral m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 895 | using neg_numeral_le_numeral [of One m] by (simp add: numeral_One) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 896 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 897 | lemma neg_numeral_less_neg_one_iff: "- numeral m < - 1 \<longleftrightarrow> m \<noteq> One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 898 | by (cases m) simp_all | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 899 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 900 | lemma neg_numeral_le_neg_one: "- numeral m \<le> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 901 | by simp | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 902 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 903 | lemma not_neg_one_less_neg_numeral: "\<not> - 1 < - numeral m" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 904 | by simp | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 905 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 906 | lemma not_neg_one_le_neg_numeral_iff: "\<not> - 1 \<le> - numeral m \<longleftrightarrow> m \<noteq> One" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 907 | by (cases m) simp_all | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 908 | |
| 47108 | 909 | lemma sub_non_negative: | 
| 910 | "sub n m \<ge> 0 \<longleftrightarrow> n \<ge> m" | |
| 911 | by (simp only: sub_def le_diff_eq) simp | |
| 912 | ||
| 913 | lemma sub_positive: | |
| 914 | "sub n m > 0 \<longleftrightarrow> n > m" | |
| 915 | by (simp only: sub_def less_diff_eq) simp | |
| 916 | ||
| 917 | lemma sub_non_positive: | |
| 918 | "sub n m \<le> 0 \<longleftrightarrow> n \<le> m" | |
| 919 | by (simp only: sub_def diff_le_eq) simp | |
| 920 | ||
| 921 | lemma sub_negative: | |
| 922 | "sub n m < 0 \<longleftrightarrow> n < m" | |
| 923 | by (simp only: sub_def diff_less_eq) simp | |
| 924 | ||
| 925 | lemmas le_neg_numeral_simps [simp] = | |
| 926 | neg_numeral_le_iff | |
| 927 | neg_numeral_le_numeral not_numeral_le_neg_numeral | |
| 928 | neg_numeral_le_zero not_zero_le_neg_numeral | |
| 929 | neg_numeral_le_one not_one_le_neg_numeral | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 930 | neg_one_le_numeral not_numeral_le_neg_one | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 931 | neg_numeral_le_neg_one not_neg_one_le_neg_numeral_iff | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 932 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 933 | lemma le_minus_one_simps [simp]: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 934 | "- 1 \<le> 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 935 | "- 1 \<le> 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 936 | "\<not> 0 \<le> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 937 | "\<not> 1 \<le> - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 938 | by simp_all | 
| 47108 | 939 | |
| 940 | lemmas less_neg_numeral_simps [simp] = | |
| 941 | neg_numeral_less_iff | |
| 942 | neg_numeral_less_numeral not_numeral_less_neg_numeral | |
| 943 | neg_numeral_less_zero not_zero_less_neg_numeral | |
| 944 | neg_numeral_less_one not_one_less_neg_numeral | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 945 | neg_one_less_numeral not_numeral_less_neg_one | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 946 | neg_numeral_less_neg_one_iff not_neg_one_less_neg_numeral | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 947 | |
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 948 | lemma less_minus_one_simps [simp]: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 949 | "- 1 < 0" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 950 | "- 1 < 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 951 | "\<not> 0 < - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 952 | "\<not> 1 < - 1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 953 | by (simp_all add: less_le) | 
| 47108 | 954 | |
| 61944 | 955 | lemma abs_numeral [simp]: "\<bar>numeral n\<bar> = numeral n" | 
| 47108 | 956 | by simp | 
| 957 | ||
| 61944 | 958 | lemma abs_neg_numeral [simp]: "\<bar>- numeral n\<bar> = numeral n" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 959 | by (simp only: abs_minus_cancel abs_numeral) | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 960 | |
| 61944 | 961 | lemma abs_neg_one [simp]: "\<bar>- 1\<bar> = 1" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 962 | by simp | 
| 47108 | 963 | |
| 964 | end | |
| 965 | ||
| 60758 | 966 | subsubsection \<open> | 
| 47108 | 967 | Natural numbers | 
| 60758 | 968 | \<close> | 
| 47108 | 969 | |
| 47299 | 970 | lemma Suc_1 [simp]: "Suc 1 = 2" | 
| 971 | unfolding Suc_eq_plus1 by (rule one_add_one) | |
| 972 | ||
| 47108 | 973 | lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" | 
| 47299 | 974 | unfolding Suc_eq_plus1 by (rule numeral_plus_one) | 
| 47108 | 975 | |
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 976 | definition pred_numeral :: "num \<Rightarrow> nat" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 977 | where [code del]: "pred_numeral k = numeral k - 1" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 978 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 979 | lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 980 | unfolding pred_numeral_def by simp | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 981 | |
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 982 | lemma eval_nat_numeral: | 
| 47108 | 983 | "numeral One = Suc 0" | 
| 984 | "numeral (Bit0 n) = Suc (numeral (BitM n))" | |
| 985 | "numeral (Bit1 n) = Suc (numeral (Bit0 n))" | |
| 986 | by (simp_all add: numeral.simps BitM_plus_one) | |
| 987 | ||
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 988 | lemma pred_numeral_simps [simp]: | 
| 47300 | 989 | "pred_numeral One = 0" | 
| 990 | "pred_numeral (Bit0 k) = numeral (BitM k)" | |
| 991 | "pred_numeral (Bit1 k) = numeral (Bit0 k)" | |
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 992 | unfolding pred_numeral_def eval_nat_numeral | 
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 993 | by (simp_all only: diff_Suc_Suc diff_0) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 994 | |
| 47192 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 995 | lemma numeral_2_eq_2: "2 = Suc (Suc 0)" | 
| 47220 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 huffman parents: 
47218diff
changeset | 996 | 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 | 997 | |
| 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 huffman parents: 
47191diff
changeset | 998 | 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 | 999 | 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 | 1000 | |
| 47207 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 1001 | 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 | 1002 | by (simp only: numeral_One One_nat_def) | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 1003 | |
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 1004 | lemma Suc_nat_number_of_add: | 
| 47300 | 1005 | "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 | 1006 | by simp | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 1007 | |
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 1008 | (*Maps #n to n for n = 1, 2*) | 
| 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 huffman parents: 
47192diff
changeset | 1009 | 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 | 1010 | |
| 60758 | 1011 | text \<open>Comparisons involving @{term Suc}.\<close>
 | 
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1012 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1013 | 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 | 1014 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1015 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1016 | 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 | 1017 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1018 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1019 | 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 | 1020 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1021 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1022 | 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 | 1023 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1024 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1025 | 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 | 1026 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1027 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1028 | 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 | 1029 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1030 | |
| 47218 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 1031 | 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 | 1032 | by (simp add: numeral_eq_Suc) | 
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 1033 | |
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 1034 | 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 | 1035 | by (simp add: numeral_eq_Suc) | 
| 
2b652cbadde1
new lemmas for simplifying subtraction on nat numerals
 huffman parents: 
47216diff
changeset | 1036 | |
| 47209 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1037 | lemma max_Suc_numeral [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1038 | "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 | 1039 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1040 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1041 | lemma max_numeral_Suc [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1042 | "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 | 1043 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1044 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1045 | lemma min_Suc_numeral [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1046 | "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 | 1047 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1048 | |
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1049 | lemma min_numeral_Suc [simp]: | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1050 | "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 | 1051 | by (simp add: numeral_eq_Suc) | 
| 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 huffman parents: 
47207diff
changeset | 1052 | |
| 60758 | 1053 | text \<open>For @{term case_nat} and @{term rec_nat}.\<close>
 | 
| 47216 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1054 | |
| 55415 | 1055 | lemma case_nat_numeral [simp]: | 
| 1056 | "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" | |
| 47216 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1057 | by (simp add: numeral_eq_Suc) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1058 | |
| 55415 | 1059 | lemma case_nat_add_eq_if [simp]: | 
| 1060 | "case_nat a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))" | |
| 47216 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1061 | by (simp add: numeral_eq_Suc) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1062 | |
| 55415 | 1063 | lemma rec_nat_numeral [simp]: | 
| 1064 | "rec_nat a f (numeral v) = | |
| 1065 | (let pv = pred_numeral v in f pv (rec_nat a f pv))" | |
| 47216 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1066 | by (simp add: numeral_eq_Suc Let_def) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1067 | |
| 55415 | 1068 | lemma rec_nat_add_eq_if [simp]: | 
| 1069 | "rec_nat a f (numeral v + n) = | |
| 1070 | (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" | |
| 47216 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1071 | by (simp add: numeral_eq_Suc Let_def) | 
| 
4d0878d54ca5
move more theorems from Nat_Numeral.thy to Num.thy
 huffman parents: 
47211diff
changeset | 1072 | |
| 60758 | 1073 | text \<open>Case analysis on @{term "n < 2"}\<close>
 | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1074 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1075 | 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 | 1076 | by (auto simp add: numeral_2_eq_2) | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1077 | |
| 60758 | 1078 | text \<open>Removal of Small Numerals: 0, 1 and (in additive positions) 2\<close> | 
| 1079 | text \<open>bh: Are these rules really a good idea?\<close> | |
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1080 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1081 | lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1082 | by simp | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1083 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1084 | lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1085 | by simp | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1086 | |
| 60758 | 1087 | text \<open>Can be used to eliminate long strings of Sucs, but not by default.\<close> | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1088 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1089 | lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1090 | by simp | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1091 | |
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47228diff
changeset | 1092 | 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 | 1093 | |
| 47108 | 1094 | |
| 60758 | 1095 | subsection \<open>Particular lemmas concerning @{term 2}\<close>
 | 
| 58512 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1096 | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59621diff
changeset | 1097 | context linordered_field | 
| 58512 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1098 | begin | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1099 | |
| 62348 | 1100 | subclass field_char_0 .. | 
| 1101 | ||
| 58512 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1102 | lemma half_gt_zero_iff: | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1103 | "0 < a / 2 \<longleftrightarrow> 0 < a" (is "?P \<longleftrightarrow> ?Q") | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1104 | by (auto simp add: field_simps) | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1105 | |
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1106 | lemma half_gt_zero [simp]: | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1107 | "0 < a \<Longrightarrow> 0 < a / 2" | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1108 | by (simp add: half_gt_zero_iff) | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1109 | |
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1110 | end | 
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1111 | |
| 
dc4d76dfa8f0
moved lemmas out of Int.thy which have nothing to do with int
 haftmann parents: 
58421diff
changeset | 1112 | |
| 60758 | 1113 | subsection \<open>Numeral equations as default simplification rules\<close> | 
| 47108 | 1114 | |
| 1115 | declare (in numeral) numeral_One [simp] | |
| 1116 | declare (in numeral) numeral_plus_numeral [simp] | |
| 1117 | declare (in numeral) add_numeral_special [simp] | |
| 1118 | declare (in neg_numeral) add_neg_numeral_simps [simp] | |
| 1119 | declare (in neg_numeral) add_neg_numeral_special [simp] | |
| 1120 | declare (in neg_numeral) diff_numeral_simps [simp] | |
| 1121 | declare (in neg_numeral) diff_numeral_special [simp] | |
| 1122 | declare (in semiring_numeral) numeral_times_numeral [simp] | |
| 1123 | declare (in ring_1) mult_neg_numeral_simps [simp] | |
| 1124 | ||
| 60758 | 1125 | subsection \<open>Setting up simprocs\<close> | 
| 47108 | 1126 | |
| 1127 | lemma mult_numeral_1: "Numeral1 * a = (a::'a::semiring_numeral)" | |
| 1128 | by simp | |
| 1129 | ||
| 1130 | lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::semiring_numeral)" | |
| 1131 | by simp | |
| 1132 | ||
| 1133 | lemma divide_numeral_1: "a / Numeral1 = (a::'a::field)" | |
| 1134 | by simp | |
| 1135 | ||
| 1136 | lemma inverse_numeral_1: | |
| 1137 | "inverse Numeral1 = (Numeral1::'a::division_ring)" | |
| 1138 | by simp | |
| 1139 | ||
| 60758 | 1140 | text\<open>Theorem lists for the cancellation simprocs. The use of a binary | 
| 1141 | numeral for 1 reduces the number of special cases.\<close> | |
| 47108 | 1142 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1143 | lemma mult_1s: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1144 | fixes a :: "'a::semiring_numeral" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1145 | and b :: "'b::ring_1" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1146 | shows "Numeral1 * a = a" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1147 | "a * Numeral1 = a" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1148 | "- Numeral1 * b = - b" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1149 | "b * - Numeral1 = - b" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1150 | by simp_all | 
| 47108 | 1151 | |
| 60758 | 1152 | setup \<open> | 
| 47226 | 1153 | Reorient_Proc.add | 
| 1154 |     (fn Const (@{const_name numeral}, _) $ _ => true
 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1155 |     | Const (@{const_name uminus}, _) $ (Const (@{const_name numeral}, _) $ _) => true
 | 
| 47226 | 1156 | | _ => false) | 
| 60758 | 1157 | \<close> | 
| 47226 | 1158 | |
| 1159 | simproc_setup reorient_numeral | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1160 |   ("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc
 | 
| 47226 | 1161 | |
| 47108 | 1162 | |
| 60758 | 1163 | subsubsection \<open>Simplification of arithmetic operations on integer constants.\<close> | 
| 47108 | 1164 | |
| 1165 | lemmas arith_special = (* already declared simp above *) | |
| 1166 | add_numeral_special add_neg_numeral_special | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1167 | diff_numeral_special | 
| 47108 | 1168 | |
| 1169 | (* rules already in simpset *) | |
| 1170 | lemmas arith_extra_simps = | |
| 1171 | numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1172 | minus_zero | 
| 47108 | 1173 | diff_numeral_simps diff_0 diff_0_right | 
| 1174 | numeral_times_numeral mult_neg_numeral_simps | |
| 1175 | mult_zero_left mult_zero_right | |
| 1176 | abs_numeral abs_neg_numeral | |
| 1177 | ||
| 60758 | 1178 | text \<open> | 
| 47108 | 1179 | For making a minimal simpset, one must include these default simprules. | 
| 61799 | 1180 | Also include \<open>simp_thms\<close>. | 
| 60758 | 1181 | \<close> | 
| 47108 | 1182 | |
| 1183 | lemmas arith_simps = | |
| 1184 | add_num_simps mult_num_simps sub_num_simps | |
| 1185 | BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps | |
| 1186 | abs_zero abs_one arith_extra_simps | |
| 1187 | ||
| 54249 | 1188 | lemmas more_arith_simps = | 
| 1189 | neg_le_iff_le | |
| 1190 | minus_zero left_minus right_minus | |
| 1191 | mult_1_left mult_1_right | |
| 1192 | mult_minus_left mult_minus_right | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 1193 | minus_add_distrib minus_minus mult.assoc | 
| 54249 | 1194 | |
| 1195 | lemmas of_nat_simps = | |
| 1196 | of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult | |
| 1197 | ||
| 60758 | 1198 | text \<open>Simplification of relational operations\<close> | 
| 47108 | 1199 | |
| 1200 | lemmas eq_numeral_extra = | |
| 1201 | zero_neq_one one_neq_zero | |
| 1202 | ||
| 1203 | lemmas rel_simps = | |
| 1204 | le_num_simps less_num_simps eq_num_simps | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1205 | le_numeral_simps le_neg_numeral_simps le_minus_one_simps le_numeral_extra | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1206 | less_numeral_simps less_neg_numeral_simps less_minus_one_simps less_numeral_extra | 
| 47108 | 1207 | eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra | 
| 1208 | ||
| 54249 | 1209 | lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)" | 
| 61799 | 1210 | \<comment> \<open>Unfold all \<open>let\<close>s involving constants\<close> | 
| 54249 | 1211 | unfolding Let_def .. | 
| 1212 | ||
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1213 | lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)" | 
| 61799 | 1214 | \<comment> \<open>Unfold all \<open>let\<close>s involving constants\<close> | 
| 54249 | 1215 | unfolding Let_def .. | 
| 1216 | ||
| 60758 | 1217 | declaration \<open> | 
| 54249 | 1218 | let | 
| 59996 | 1219 | fun number_of ctxt T n = | 
| 1220 |     if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort numeral}))
 | |
| 54249 | 1221 |     then raise CTERM ("number_of", [])
 | 
| 59996 | 1222 | else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; | 
| 54249 | 1223 | in | 
| 1224 | K ( | |
| 1225 |     Lin_Arith.add_simps (@{thms arith_simps} @ @{thms more_arith_simps}
 | |
| 1226 |       @ @{thms rel_simps}
 | |
| 1227 |       @ @{thms pred_numeral_simps}
 | |
| 1228 |       @ @{thms arith_special numeral_One}
 | |
| 1229 |       @ @{thms of_nat_simps})
 | |
| 1230 |     #> Lin_Arith.add_simps [@{thm Suc_numeral},
 | |
| 1231 |       @{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1},
 | |
| 1232 |       @{thm le_Suc_numeral}, @{thm le_numeral_Suc},
 | |
| 1233 |       @{thm less_Suc_numeral}, @{thm less_numeral_Suc},
 | |
| 1234 |       @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc},
 | |
| 1235 |       @{thm mult_Suc}, @{thm mult_Suc_right},
 | |
| 1236 |       @{thm of_nat_numeral}]
 | |
| 1237 | #> Lin_Arith.set_number_of number_of) | |
| 1238 | end | |
| 60758 | 1239 | \<close> | 
| 54249 | 1240 | |
| 47108 | 1241 | |
| 60758 | 1242 | subsubsection \<open>Simplification of arithmetic when nested to the right.\<close> | 
| 47108 | 1243 | |
| 1244 | lemma add_numeral_left [simp]: | |
| 1245 | "numeral v + (numeral w + z) = (numeral(v + w) + z)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 1246 | by (simp_all add: add.assoc [symmetric]) | 
| 47108 | 1247 | |
| 1248 | lemma add_neg_numeral_left [simp]: | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1249 | "numeral v + (- numeral w + y) = (sub v w + y)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1250 | "- numeral v + (numeral w + y) = (sub w v + y)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1251 | "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 1252 | by (simp_all add: add.assoc [symmetric]) | 
| 47108 | 1253 | |
| 1254 | lemma mult_numeral_left [simp]: | |
| 1255 | "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1256 | "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'b::ring_1)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1257 | "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'b::ring_1)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 1258 | "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55974diff
changeset | 1259 | by (simp_all add: mult.assoc [symmetric]) | 
| 47108 | 1260 | |
| 1261 | hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec | |
| 1262 | ||
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50817diff
changeset | 1263 | |
| 60758 | 1264 | subsection \<open>code module namespace\<close> | 
| 47108 | 1265 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52210diff
changeset | 1266 | code_identifier | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52210diff
changeset | 1267 | code_module Num \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | 
| 47108 | 1268 | |
| 1269 | end |