src/HOL/Nat.ML
author paulson
Mon, 22 Oct 2001 11:54:22 +0200
changeset 11868 56db9f3a6b3e
parent 11701 3d51fbf81c17
child 12486 0ed8bdd883e0
permissions -rw-r--r--
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
to their abstract counterparts, while other binary numerals work correctly.
     1 (*  Title:      HOL/Nat.ML
     2     ID:         $Id$
     3     Author:     Lawrence C Paulson and Tobias Nipkow
     4 
     5 Proofs about natural numbers and elementary arithmetic: addition,
     6 multiplication, etc.  Some from the Hoare example from Norbert Galm.
     7 *)
     8 
     9 (** conversion rules for nat_rec **)
    10 
    11 val [nat_rec_0, nat_rec_Suc] = nat.recs;
    12 bind_thm ("nat_rec_0", nat_rec_0);
    13 bind_thm ("nat_rec_Suc", nat_rec_Suc);
    14 
    15 (*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
    16 val prems = Goal
    17     "[| !!n. f(n) == nat_rec c h n |] ==> f(0) = c";
    18 by (simp_tac (simpset() addsimps prems) 1);
    19 qed "def_nat_rec_0";
    20 
    21 val prems = Goal
    22     "[| !!n. f(n) == nat_rec c h n |] ==> f(Suc(n)) = h n (f n)";
    23 by (simp_tac (simpset() addsimps prems) 1);
    24 qed "def_nat_rec_Suc";
    25 
    26 val [nat_case_0, nat_case_Suc] = nat.cases;
    27 bind_thm ("nat_case_0", nat_case_0);
    28 bind_thm ("nat_case_Suc", nat_case_Suc);
    29 
    30 Goal "n ~= 0 ==> EX m. n = Suc m";
    31 by (case_tac "n" 1);
    32 by (REPEAT (Blast_tac 1));
    33 qed "not0_implies_Suc";
    34 
    35 Goal "!!n::nat. m<n ==> n ~= 0";
    36 by (case_tac "n" 1);
    37 by (ALLGOALS Asm_full_simp_tac);
    38 qed "gr_implies_not0";
    39 
    40 Goal "!!n::nat. (n ~= 0) = (0 < n)";
    41 by (case_tac "n" 1);
    42 by Auto_tac;
    43 qed "neq0_conv";
    44 AddIffs [neq0_conv];
    45 
    46 (*This theorem is useful with blast_tac: (n=0 ==> False) ==> 0<n *)
    47 bind_thm ("gr0I", [neq0_conv, notI] MRS iffD1);
    48 
    49 Goal "(0<n) = (EX m. n = Suc m)";
    50 by(fast_tac (claset() addIs [not0_implies_Suc]) 1);
    51 qed "gr0_conv_Suc";
    52 
    53 Goal "!!n::nat. (~(0 < n)) = (n=0)";
    54 by (rtac iffI 1);
    55  by (rtac ccontr 1);
    56  by (ALLGOALS Asm_full_simp_tac);
    57 qed "not_gr0";
    58 AddIffs [not_gr0];
    59 
    60 Goal "(Suc n <= m') --> (? m. m' = Suc m)";
    61 by (induct_tac "m'" 1);
    62 by  Auto_tac;
    63 qed_spec_mp "Suc_le_D";
    64 
    65 (*Useful in certain inductive arguments*)
    66 Goal "(m < Suc n) = (m=0 | (EX j. m = Suc j & j < n))";
    67 by (case_tac "m" 1);
    68 by Auto_tac;
    69 qed "less_Suc_eq_0_disj";
    70 
    71 val prems = Goal "[| P 0; P(Suc 0); !!k. P k ==> P (Suc (Suc k)) |] ==> P n";
    72 by (rtac nat_less_induct 1);
    73 by (case_tac "n" 1);
    74 by (case_tac "nat" 2);
    75 by (ALLGOALS (blast_tac (claset() addIs prems@[less_trans])));
    76 qed "nat_induct2";
    77 
    78 (** LEAST theorems for type "nat" by specialization **)
    79 
    80 bind_thm("LeastI", wellorder_LeastI);
    81 bind_thm("Least_le", wellorder_Least_le);
    82 bind_thm("not_less_Least", wellorder_not_less_Least);
    83 
    84 Goal "[| P n; ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))";
    85 by (case_tac "n" 1);
    86 by Auto_tac;  
    87 by (ftac LeastI 1); 
    88 by (dres_inst_tac [("P","%x. P (Suc x)")] LeastI 1);
    89 by (subgoal_tac "(LEAST x. P x) <= Suc (LEAST x. P (Suc x))" 1); 
    90 by (etac Least_le 2); 
    91 by (case_tac "LEAST x. P x" 1);
    92 by Auto_tac;  
    93 by (dres_inst_tac [("P","%x. P (Suc x)")] Least_le 1);
    94 by (blast_tac (claset() addIs [order_antisym]) 1); 
    95 qed "Least_Suc";
    96 
    97 Goal "[|P n; Q m; ~P 0; !k. P (Suc k) = Q k|] ==> Least P = Suc (Least Q)";
    98 by (eatac (Least_Suc RS ssubst) 1 1);
    99 by (Asm_simp_tac 1);
   100 qed "Least_Suc2";
   101 
   102 
   103 (** min and max **)
   104 
   105 Goal "min 0 n = (0::nat)";
   106 by (rtac min_leastL 1);
   107 by (Simp_tac 1);
   108 qed "min_0L";
   109 
   110 Goal "min n 0 = (0::nat)";
   111 by (rtac min_leastR 1);
   112 by (Simp_tac 1);
   113 qed "min_0R";
   114 
   115 Goal "min (Suc m) (Suc n) = Suc (min m n)";
   116 by (simp_tac (simpset() addsimps [min_of_mono]) 1);
   117 qed "min_Suc_Suc";
   118 
   119 Addsimps [min_0L,min_0R,min_Suc_Suc];
   120 
   121 Goal "max 0 n = (n::nat)";
   122 by (rtac max_leastL 1);
   123 by (Simp_tac 1);
   124 qed "max_0L";
   125 
   126 Goal "max n 0 = (n::nat)";
   127 by (rtac max_leastR 1);
   128 by (Simp_tac 1);
   129 qed "max_0R";
   130 
   131 Goal "max (Suc m) (Suc n) = Suc(max m n)";
   132 by (simp_tac (simpset() addsimps [max_of_mono]) 1);
   133 qed "max_Suc_Suc";
   134 
   135 Addsimps [max_0L,max_0R,max_Suc_Suc];
   136 
   137 
   138 (*** Basic rewrite rules for the arithmetic operators ***)
   139 
   140 (** Difference **)
   141 
   142 Goal "0 - n = (0::nat)";
   143 by (induct_tac "n" 1);
   144 by (ALLGOALS Asm_simp_tac);
   145 qed "diff_0_eq_0";
   146 
   147 (*Must simplify BEFORE the induction!  (Else we get a critical pair)
   148   Suc(m) - Suc(n)   rewrites to   pred(Suc(m) - n)  *)
   149 Goal "Suc(m) - Suc(n) = m - n";
   150 by (Simp_tac 1);
   151 by (induct_tac "n" 1);
   152 by (ALLGOALS Asm_simp_tac);
   153 qed "diff_Suc_Suc";
   154 
   155 Addsimps [diff_0_eq_0, diff_Suc_Suc];
   156 
   157 (* Could be (and is, below) generalized in various ways;
   158    However, none of the generalizations are currently in the simpset,
   159    and I dread to think what happens if I put them in *)
   160 Goal "0 < n ==> Suc(n - Suc 0) = n";
   161 by (asm_simp_tac (simpset() addsplits [nat.split]) 1);
   162 qed "Suc_pred";
   163 Addsimps [Suc_pred];
   164 
   165 Delsimps [diff_Suc];
   166 
   167 
   168 (**** Inductive properties of the operators ****)
   169 
   170 (*** Addition ***)
   171 
   172 Goal "m + 0 = (m::nat)";
   173 by (induct_tac "m" 1);
   174 by (ALLGOALS Asm_simp_tac);
   175 qed "add_0_right";
   176 
   177 Goal "m + Suc(n) = Suc(m+n)";
   178 by (induct_tac "m" 1);
   179 by (ALLGOALS Asm_simp_tac);
   180 qed "add_Suc_right";
   181 
   182 Addsimps [add_0_right,add_Suc_right];
   183 
   184 
   185 (*Associative law for addition*)
   186 Goal "(m + n) + k = m + ((n + k)::nat)";
   187 by (induct_tac "m" 1);
   188 by (ALLGOALS Asm_simp_tac);
   189 qed "add_assoc";
   190 
   191 (*Commutative law for addition*)
   192 Goal "m + n = n + (m::nat)";
   193 by (induct_tac "m" 1);
   194 by (ALLGOALS Asm_simp_tac);
   195 qed "add_commute";
   196 
   197 Goal "x+(y+z)=y+((x+z)::nat)";
   198 by (rtac (add_commute RS trans) 1);
   199 by (rtac (add_assoc RS trans) 1);
   200 by (rtac (add_commute RS arg_cong) 1);
   201 qed "add_left_commute";
   202 
   203 (*Addition is an AC-operator*)
   204 bind_thms ("add_ac", [add_assoc, add_commute, add_left_commute]);
   205 
   206 Goal "(k + m = k + n) = (m=(n::nat))";
   207 by (induct_tac "k" 1);
   208 by (Simp_tac 1);
   209 by (Asm_simp_tac 1);
   210 qed "add_left_cancel";
   211 
   212 Goal "(m + k = n + k) = (m=(n::nat))";
   213 by (induct_tac "k" 1);
   214 by (Simp_tac 1);
   215 by (Asm_simp_tac 1);
   216 qed "add_right_cancel";
   217 
   218 Goal "(k + m <= k + n) = (m<=(n::nat))";
   219 by (induct_tac "k" 1);
   220 by (Simp_tac 1);
   221 by (Asm_simp_tac 1);
   222 qed "add_left_cancel_le";
   223 
   224 Goal "(k + m < k + n) = (m<(n::nat))";
   225 by (induct_tac "k" 1);
   226 by (Simp_tac 1);
   227 by (Asm_simp_tac 1);
   228 qed "add_left_cancel_less";
   229 
   230 Addsimps [add_left_cancel, add_right_cancel,
   231           add_left_cancel_le, add_left_cancel_less];
   232 
   233 (** Reasoning about m+0=0, etc. **)
   234 
   235 Goal "!!m::nat. (m+n = 0) = (m=0 & n=0)";
   236 by (case_tac "m" 1);
   237 by (Auto_tac);
   238 qed "add_is_0";
   239 AddIffs [add_is_0];
   240 
   241 Goal "(m+n= Suc 0) = (m= Suc 0 & n=0 | m=0 & n= Suc 0)";
   242 by (case_tac "m" 1);
   243 by (Auto_tac);
   244 qed "add_is_1";
   245 
   246 Goal "(Suc 0 = m+n) = (m = Suc 0 & n=0 | m=0 & n = Suc 0)";
   247 by (rtac ([eq_commute, add_is_1] MRS trans) 1);
   248 qed "one_is_add";
   249 
   250 Goal "!!m::nat. (0<m+n) = (0<m | 0<n)";
   251 by (simp_tac (simpset() delsimps [neq0_conv] addsimps [neq0_conv RS sym]) 1);
   252 qed "add_gr_0";
   253 AddIffs [add_gr_0];
   254 
   255 Goal "!!m::nat. m + n = m ==> n = 0";
   256 by (dtac (add_0_right RS ssubst) 1);
   257 by (asm_full_simp_tac (simpset() addsimps [add_assoc]
   258                                  delsimps [add_0_right]) 1);
   259 qed "add_eq_self_zero";
   260 
   261 
   262 (**** Additional theorems about "less than" ****)
   263 
   264 (*Deleted less_natE; instead use less_imp_Suc_add RS exE*)
   265 Goal "m<n --> (EX k. n=Suc(m+k))";
   266 by (induct_tac "n" 1);
   267 by (ALLGOALS (simp_tac (simpset() addsimps [order_le_less])));
   268 by (blast_tac (claset() addSEs [less_SucE]
   269                         addSIs [add_0_right RS sym, add_Suc_right RS sym]) 1);
   270 qed_spec_mp "less_imp_Suc_add";
   271 
   272 Goal "n <= ((m + n)::nat)";
   273 by (induct_tac "m" 1);
   274 by (ALLGOALS Simp_tac);
   275 by (etac le_SucI 1);
   276 qed "le_add2";
   277 
   278 Goal "n <= ((n + m)::nat)";
   279 by (simp_tac (simpset() addsimps add_ac) 1);
   280 by (rtac le_add2 1);
   281 qed "le_add1";
   282 
   283 bind_thm ("less_add_Suc1", (lessI RS (le_add1 RS le_less_trans)));
   284 bind_thm ("less_add_Suc2", (lessI RS (le_add2 RS le_less_trans)));
   285 
   286 Goal "(m<n) = (EX k. n=Suc(m+k))";
   287 by (blast_tac (claset() addSIs [less_add_Suc1, less_imp_Suc_add]) 1);
   288 qed "less_iff_Suc_add";
   289 
   290 
   291 (*"i <= j ==> i <= j+m"*)
   292 bind_thm ("trans_le_add1", le_add1 RSN (2,le_trans));
   293 
   294 (*"i <= j ==> i <= m+j"*)
   295 bind_thm ("trans_le_add2", le_add2 RSN (2,le_trans));
   296 
   297 (*"i < j ==> i < j+m"*)
   298 bind_thm ("trans_less_add1", le_add1 RSN (2,less_le_trans));
   299 
   300 (*"i < j ==> i < m+j"*)
   301 bind_thm ("trans_less_add2", le_add2 RSN (2,less_le_trans));
   302 
   303 Goal "i+j < (k::nat) --> i<k";
   304 by (induct_tac "j" 1);
   305 by (ALLGOALS Asm_simp_tac);
   306 by (blast_tac (claset() addDs [Suc_lessD]) 1);
   307 qed_spec_mp "add_lessD1";
   308 
   309 Goal "~ (i+j < (i::nat))";
   310 by (rtac notI 1);
   311 by (etac (add_lessD1 RS less_irrefl) 1);
   312 qed "not_add_less1";
   313 
   314 Goal "~ (j+i < (i::nat))";
   315 by (simp_tac (simpset() addsimps [add_commute, not_add_less1]) 1);
   316 qed "not_add_less2";
   317 AddIffs [not_add_less1, not_add_less2];
   318 
   319 Goal "m+k<=n --> m<=(n::nat)";
   320 by (induct_tac "k" 1);
   321 by (ALLGOALS (asm_simp_tac (simpset() addsimps le_simps)));
   322 qed_spec_mp "add_leD1";
   323 
   324 Goal "m+k<=n ==> k<=(n::nat)";
   325 by (full_simp_tac (simpset() addsimps [add_commute]) 1);
   326 by (etac add_leD1 1);
   327 qed_spec_mp "add_leD2";
   328 
   329 Goal "m+k<=n ==> m<=n & k<=(n::nat)";
   330 by (blast_tac (claset() addDs [add_leD1, add_leD2]) 1);
   331 bind_thm ("add_leE", result() RS conjE);
   332 
   333 (*needs !!k for add_ac to work*)
   334 Goal "!!k:: nat. [| k<l;  m+l = k+n |] ==> m<n";
   335 by (force_tac (claset(),
   336               simpset() delsimps [add_Suc_right]
   337                         addsimps [less_iff_Suc_add,
   338                                   add_Suc_right RS sym] @ add_ac) 1);
   339 qed "less_add_eq_less";
   340 
   341 
   342 (*** Monotonicity of Addition ***)
   343 
   344 (*strict, in 1st argument*)
   345 Goal "i < j ==> i + k < j + (k::nat)";
   346 by (induct_tac "k" 1);
   347 by (ALLGOALS Asm_simp_tac);
   348 qed "add_less_mono1";
   349 
   350 (*strict, in both arguments*)
   351 Goal "[|i < j; k < l|] ==> i + k < j + (l::nat)";
   352 by (rtac (add_less_mono1 RS less_trans) 1);
   353 by (REPEAT (assume_tac 1));
   354 by (induct_tac "j" 1);
   355 by (ALLGOALS Asm_simp_tac);
   356 qed "add_less_mono";
   357 
   358 (*A [clumsy] way of lifting < monotonicity to <= monotonicity *)
   359 val [lt_mono,le] = Goal
   360      "[| !!i j::nat. i<j ==> f(i) < f(j);       \
   361 \        i <= j                                 \
   362 \     |] ==> f(i) <= (f(j)::nat)";
   363 by (cut_facts_tac [le] 1);
   364 by (asm_full_simp_tac (simpset() addsimps [order_le_less]) 1);
   365 by (blast_tac (claset() addSIs [lt_mono]) 1);
   366 qed "less_mono_imp_le_mono";
   367 
   368 (*non-strict, in 1st argument*)
   369 Goal "i<=j ==> i + k <= j + (k::nat)";
   370 by (res_inst_tac [("f", "%j. j+k")] less_mono_imp_le_mono 1);
   371 by (etac add_less_mono1 1);
   372 by (assume_tac 1);
   373 qed "add_le_mono1";
   374 
   375 (*non-strict, in both arguments*)
   376 Goal "[|i<=j;  k<=l |] ==> i + k <= j + (l::nat)";
   377 by (etac (add_le_mono1 RS le_trans) 1);
   378 by (simp_tac (simpset() addsimps [add_commute]) 1);
   379 qed "add_le_mono";
   380 
   381 
   382 (*** Multiplication ***)
   383 
   384 (*right annihilation in product*)
   385 Goal "!!m::nat. m * 0 = 0";
   386 by (induct_tac "m" 1);
   387 by (ALLGOALS Asm_simp_tac);
   388 qed "mult_0_right";
   389 
   390 (*right successor law for multiplication*)
   391 Goal  "m * Suc(n) = m + (m * n)";
   392 by (induct_tac "m" 1);
   393 by (ALLGOALS(asm_simp_tac (simpset() addsimps add_ac)));
   394 qed "mult_Suc_right";
   395 
   396 Addsimps [mult_0_right, mult_Suc_right];
   397 
   398 Goal "(1::nat) * n = n";
   399 by (Asm_simp_tac 1);
   400 qed "mult_1";
   401 
   402 Goal "n * (1::nat) = n";
   403 by (Asm_simp_tac 1);
   404 qed "mult_1_right";
   405 
   406 (*Commutative law for multiplication*)
   407 Goal "m * n = n * (m::nat)";
   408 by (induct_tac "m" 1);
   409 by (ALLGOALS Asm_simp_tac);
   410 qed "mult_commute";
   411 
   412 (*addition distributes over multiplication*)
   413 Goal "(m + n)*k = (m*k) + ((n*k)::nat)";
   414 by (induct_tac "m" 1);
   415 by (ALLGOALS(asm_simp_tac (simpset() addsimps add_ac)));
   416 qed "add_mult_distrib";
   417 
   418 Goal "k*(m + n) = (k*m) + ((k*n)::nat)";
   419 by (induct_tac "m" 1);
   420 by (ALLGOALS(asm_simp_tac (simpset() addsimps add_ac)));
   421 qed "add_mult_distrib2";
   422 
   423 (*Associative law for multiplication*)
   424 Goal "(m * n) * k = m * ((n * k)::nat)";
   425 by (induct_tac "m" 1);
   426 by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_mult_distrib])));
   427 qed "mult_assoc";
   428 
   429 Goal "x*(y*z) = y*((x*z)::nat)";
   430 by (rtac trans 1);
   431 by (rtac mult_commute 1);
   432 by (rtac trans 1);
   433 by (rtac mult_assoc 1);
   434 by (rtac (mult_commute RS arg_cong) 1);
   435 qed "mult_left_commute";
   436 
   437 bind_thms ("mult_ac", [mult_assoc,mult_commute,mult_left_commute]);
   438 
   439 Goal "!!m::nat. (m*n = 0) = (m=0 | n=0)";
   440 by (induct_tac "m" 1);
   441 by (induct_tac "n" 2);
   442 by (ALLGOALS Asm_simp_tac);
   443 qed "mult_is_0";
   444 Addsimps [mult_is_0];
   445 
   446 
   447 (*** Difference ***)
   448 
   449 Goal "!!m::nat. m - m = 0";
   450 by (induct_tac "m" 1);
   451 by (ALLGOALS Asm_simp_tac);
   452 qed "diff_self_eq_0";
   453 
   454 Addsimps [diff_self_eq_0];
   455 
   456 (*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *)
   457 Goal "~ m<n --> n+(m-n) = (m::nat)";
   458 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
   459 by (ALLGOALS Asm_simp_tac);
   460 qed_spec_mp "add_diff_inverse";
   461 
   462 Goal "n<=m ==> n+(m-n) = (m::nat)";
   463 by (asm_simp_tac (simpset() addsimps [add_diff_inverse, not_less_iff_le]) 1);
   464 qed "le_add_diff_inverse";
   465 
   466 Goal "n<=m ==> (m-n)+n = (m::nat)";
   467 by (asm_simp_tac (simpset() addsimps [le_add_diff_inverse, add_commute]) 1);
   468 qed "le_add_diff_inverse2";
   469 
   470 Addsimps  [le_add_diff_inverse, le_add_diff_inverse2];
   471 
   472 
   473 (*** More results about difference ***)
   474 
   475 Goal "n <= m ==> Suc(m)-n = Suc(m-n)";
   476 by (etac rev_mp 1);
   477 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
   478 by (ALLGOALS Asm_simp_tac);
   479 qed "Suc_diff_le";
   480 
   481 Goal "m - n < Suc(m)";
   482 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
   483 by (etac less_SucE 3);
   484 by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq])));
   485 qed "diff_less_Suc";
   486 
   487 Goal "m - n <= (m::nat)";
   488 by (res_inst_tac [("m","m"), ("n","n")] diff_induct 1);
   489 by (ALLGOALS (asm_simp_tac (simpset() addsimps [le_SucI])));
   490 qed "diff_le_self";
   491 Addsimps [diff_le_self];
   492 
   493 (* j<k ==> j-n < k *)
   494 bind_thm ("less_imp_diff_less", diff_le_self RS le_less_trans);
   495 
   496 Goal "!!i::nat. i-j-k = i - (j+k)";
   497 by (res_inst_tac [("m","i"),("n","j")] diff_induct 1);
   498 by (ALLGOALS Asm_simp_tac);
   499 qed "diff_diff_left";
   500 
   501 Goal "(Suc m - n) - Suc k = m - n - k";
   502 by (simp_tac (simpset() addsimps [diff_diff_left]) 1);
   503 qed "Suc_diff_diff";
   504 Addsimps [Suc_diff_diff];
   505 
   506 Goal "0<n ==> n - Suc i < n";
   507 by (case_tac "n" 1);
   508 by Safe_tac;
   509 by (asm_simp_tac (simpset() addsimps le_simps) 1);
   510 qed "diff_Suc_less";
   511 Addsimps [diff_Suc_less];
   512 
   513 (*This and the next few suggested by Florian Kammueller*)
   514 Goal "!!i::nat. i-j-k = i-k-j";
   515 by (simp_tac (simpset() addsimps [diff_diff_left, add_commute]) 1);
   516 qed "diff_commute";
   517 
   518 Goal "k <= (j::nat) --> (i + j) - k = i + (j - k)";
   519 by (res_inst_tac [("m","j"),("n","k")] diff_induct 1);
   520 by (ALLGOALS Asm_simp_tac);
   521 qed_spec_mp "diff_add_assoc";
   522 
   523 Goal "k <= (j::nat) --> (j + i) - k = (j - k) + i";
   524 by (asm_simp_tac (simpset() addsimps [add_commute, diff_add_assoc]) 1);
   525 qed_spec_mp "diff_add_assoc2";
   526 
   527 Goal "(n+m) - n = (m::nat)";
   528 by (induct_tac "n" 1);
   529 by (ALLGOALS Asm_simp_tac);
   530 qed "diff_add_inverse";
   531 
   532 Goal "(m+n) - n = (m::nat)";
   533 by (simp_tac (simpset() addsimps [diff_add_assoc]) 1);
   534 qed "diff_add_inverse2";
   535 
   536 Goal "i <= (j::nat) ==> (j-i=k) = (j=k+i)";
   537 by Safe_tac;
   538 by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_add_inverse2])));
   539 qed "le_imp_diff_is_add";
   540 
   541 Goal "!!m::nat. (m-n = 0) = (m <= n)";
   542 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
   543 by (ALLGOALS Asm_simp_tac);
   544 qed "diff_is_0_eq";
   545 Addsimps [diff_is_0_eq];
   546 
   547 Goal "!!m::nat. (0<n-m) = (m<n)";
   548 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
   549 by (ALLGOALS Asm_simp_tac);
   550 qed "zero_less_diff";
   551 Addsimps [zero_less_diff];
   552 
   553 Goal "i < j  ==> EX k::nat. 0<k & i+k = j";
   554 by (res_inst_tac [("x","j - i")] exI 1);
   555 by (asm_simp_tac (simpset() addsimps [add_diff_inverse, less_not_sym]) 1);
   556 qed "less_imp_add_positive";
   557 
   558 Goal "P(k) --> (ALL n. P(Suc(n))--> P(n)) --> P(k-i)";
   559 by (res_inst_tac [("m","k"),("n","i")] diff_induct 1);
   560 by (ALLGOALS (Clarify_tac THEN' Simp_tac THEN' TRY o Blast_tac));
   561 qed "zero_induct_lemma";
   562 
   563 val prems = Goal "[| P(k);  !!n. P(Suc(n)) ==> P(n) |] ==> P(0)";
   564 by (rtac (diff_self_eq_0 RS subst) 1);
   565 by (rtac (zero_induct_lemma RS mp RS mp) 1);
   566 by (REPEAT (ares_tac ([impI,allI]@prems) 1));
   567 qed "zero_induct";
   568 
   569 Goal "(k+m) - (k+n) = m - (n::nat)";
   570 by (induct_tac "k" 1);
   571 by (ALLGOALS Asm_simp_tac);
   572 qed "diff_cancel";
   573 
   574 Goal "(m+k) - (n+k) = m - (n::nat)";
   575 by (asm_simp_tac
   576     (simpset() addsimps [diff_cancel, inst "n" "k" add_commute]) 1);
   577 qed "diff_cancel2";
   578 
   579 Goal "n - (n+m) = (0::nat)";
   580 by (induct_tac "n" 1);
   581 by (ALLGOALS Asm_simp_tac);
   582 qed "diff_add_0";
   583 
   584 
   585 (** Difference distributes over multiplication **)
   586 
   587 Goal "!!m::nat. (m - n) * k = (m * k) - (n * k)";
   588 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
   589 by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_cancel])));
   590 qed "diff_mult_distrib" ;
   591 
   592 Goal "!!m::nat. k * (m - n) = (k * m) - (k * n)";
   593 val mult_commute_k = read_instantiate [("m","k")] mult_commute;
   594 by (simp_tac (simpset() addsimps [diff_mult_distrib, mult_commute_k]) 1);
   595 qed "diff_mult_distrib2" ;
   596 (*NOT added as rewrites, since sometimes they are used from right-to-left*)
   597 
   598 bind_thms ("nat_distrib",
   599   [add_mult_distrib, add_mult_distrib2, diff_mult_distrib, diff_mult_distrib2]);
   600 
   601 
   602 (*** Monotonicity of Multiplication ***)
   603 
   604 Goal "i <= (j::nat) ==> i*k<=j*k";
   605 by (induct_tac "k" 1);
   606 by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono])));
   607 qed "mult_le_mono1";
   608 
   609 Goal "i <= (j::nat) ==> k*i <= k*j";
   610 by (dtac mult_le_mono1 1);
   611 by (asm_simp_tac (simpset() addsimps [mult_commute]) 1);
   612 qed "mult_le_mono2";
   613 
   614 (* <= monotonicity, BOTH arguments*)
   615 Goal "[| i <= (j::nat); k <= l |] ==> i*k <= j*l";
   616 by (etac (mult_le_mono1 RS le_trans) 1);
   617 by (etac mult_le_mono2 1);
   618 qed "mult_le_mono";
   619 
   620 (*strict, in 1st argument; proof is by induction on k>0*)
   621 Goal "!!i::nat. [| i<j; 0<k |] ==> k*i < k*j";
   622 by (eres_inst_tac [("m1","0")] (less_imp_Suc_add RS exE) 1);
   623 by (Asm_simp_tac 1);
   624 by (induct_tac "x" 1);
   625 by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_less_mono])));
   626 qed "mult_less_mono2";
   627 
   628 Goal "!!i::nat. [| i<j; 0<k |] ==> i*k < j*k";
   629 by (dtac mult_less_mono2 1);
   630 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [mult_commute])));
   631 qed "mult_less_mono1";
   632 
   633 Goal "!!m::nat. (0 < m*n) = (0<m & 0<n)";
   634 by (induct_tac "m" 1);
   635 by (case_tac "n" 2);
   636 by (ALLGOALS Asm_simp_tac);
   637 qed "zero_less_mult_iff";
   638 Addsimps [zero_less_mult_iff];
   639 
   640 Goal "(Suc 0 <= m*n) = (1<=m & 1<=n)";
   641 by (induct_tac "m" 1);
   642 by (case_tac "n" 2);
   643 by (ALLGOALS Asm_simp_tac);
   644 qed "one_le_mult_iff";
   645 Addsimps [one_le_mult_iff];
   646 
   647 Goal "(m*n = Suc 0) = (m=1 & n=1)";
   648 by (induct_tac "m" 1);
   649 by (Simp_tac 1);
   650 by (induct_tac "n" 1);
   651 by (Simp_tac 1);
   652 by (fast_tac (claset() addss simpset()) 1);
   653 qed "mult_eq_1_iff";
   654 Addsimps [mult_eq_1_iff];
   655 
   656 Goal "(Suc 0 = m*n) = (m=1 & n=1)";
   657 by(rtac (mult_eq_1_iff RSN (2,trans)) 1);
   658 by (fast_tac (claset() addss simpset()) 1);
   659 qed "one_eq_mult_iff";
   660 Addsimps [one_eq_mult_iff];
   661 
   662 Goal "!!m::nat. (m*k < n*k) = (0<k & m<n)";
   663 by (safe_tac (claset() addSIs [mult_less_mono1]));
   664 by (case_tac "k" 1);
   665 by Auto_tac;  
   666 by (full_simp_tac (simpset() delsimps [le_0_eq]
   667 			     addsimps [linorder_not_le RS sym]) 1);
   668 by (blast_tac (claset() addIs [mult_le_mono1]) 1); 
   669 qed "mult_less_cancel2";
   670 
   671 Goal "!!m::nat. (k*m < k*n) = (0<k & m<n)";
   672 by (simp_tac (simpset() addsimps [mult_less_cancel2, 
   673                                   inst "m" "k" mult_commute]) 1);
   674 qed "mult_less_cancel1";
   675 Addsimps [mult_less_cancel1, mult_less_cancel2];
   676 
   677 Goal "!!m::nat. (m*k <= n*k) = (0<k --> m<=n)";
   678 by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
   679 by Auto_tac;  
   680 qed "mult_le_cancel2";
   681 
   682 Goal "!!m::nat. (k*m <= k*n) = (0<k --> m<=n)";
   683 by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
   684 by Auto_tac;  
   685 qed "mult_le_cancel1";
   686 Addsimps [mult_le_cancel1, mult_le_cancel2];
   687 
   688 Goal "(m*k = n*k) = (m=n | (k = (0::nat)))";
   689 by (cut_facts_tac [less_linear] 1);
   690 by Safe_tac;
   691 by Auto_tac; 	
   692 by (ALLGOALS (dtac mult_less_mono1 THEN' assume_tac));
   693 by (ALLGOALS Asm_full_simp_tac);
   694 qed "mult_cancel2";
   695 
   696 Goal "(k*m = k*n) = (m=n | (k = (0::nat)))";
   697 by (simp_tac (simpset() addsimps [mult_cancel2, inst "m" "k" mult_commute]) 1);
   698 qed "mult_cancel1";
   699 Addsimps [mult_cancel1, mult_cancel2];
   700 
   701 Goal "(Suc k * m < Suc k * n) = (m < n)";
   702 by (stac mult_less_cancel1 1);
   703 by (Simp_tac 1);
   704 qed "Suc_mult_less_cancel1";
   705 
   706 Goal "(Suc k * m <= Suc k * n) = (m <= n)";
   707 by (stac mult_le_cancel1 1);
   708 by (Simp_tac 1);
   709 qed "Suc_mult_le_cancel1";
   710 
   711 Goal "(Suc k * m = Suc k * n) = (m = n)";
   712 by (stac mult_cancel1 1);
   713 by (Simp_tac 1);
   714 qed "Suc_mult_cancel1";
   715 
   716 
   717 (*Lemma for gcd*)
   718 Goal "!!m::nat. m = m*n ==> n=1 | m=0";
   719 by (dtac sym 1);
   720 by (rtac disjCI 1);
   721 by (rtac nat_less_cases 1 THEN assume_tac 2);
   722 by (fast_tac (claset() addSEs [less_SucE] addss simpset()) 1);
   723 by (best_tac (claset() addDs [mult_less_mono2] addss simpset()) 1);
   724 qed "mult_eq_self_implies_10";