377 by (etac rev_mp 1); |
377 by (etac rev_mp 1); |
378 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
378 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
379 by (ALLGOALS Asm_simp_tac); |
379 by (ALLGOALS Asm_simp_tac); |
380 qed "Suc_diff_le"; |
380 qed "Suc_diff_le"; |
381 |
381 |
382 Goal "n<=(l::nat) --> Suc l - n + m = Suc (l - n + m)"; |
|
383 by (res_inst_tac [("m","n"),("n","l")] diff_induct 1); |
|
384 by (ALLGOALS Asm_simp_tac); |
|
385 qed_spec_mp "Suc_diff_add_le"; |
|
386 |
|
387 Goal "m - n < Suc(m)"; |
382 Goal "m - n < Suc(m)"; |
388 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
383 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
389 by (etac less_SucE 3); |
384 by (etac less_SucE 3); |
390 by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq]))); |
385 by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq]))); |
391 qed "diff_less_Suc"; |
386 qed "diff_less_Suc"; |
414 by Safe_tac; |
409 by Safe_tac; |
415 by (asm_simp_tac (simpset() addsimps le_simps) 1); |
410 by (asm_simp_tac (simpset() addsimps le_simps) 1); |
416 qed "diff_Suc_less"; |
411 qed "diff_Suc_less"; |
417 Addsimps [diff_Suc_less]; |
412 Addsimps [diff_Suc_less]; |
418 |
413 |
419 Goal "i<n ==> n - Suc i < n - i"; |
|
420 by (exhaust_tac "n" 1); |
|
421 by (auto_tac (claset(), |
|
422 simpset() addsimps [Suc_diff_le]@le_simps)); |
|
423 qed "diff_Suc_less_diff"; |
|
424 |
|
425 (*This and the next few suggested by Florian Kammueller*) |
414 (*This and the next few suggested by Florian Kammueller*) |
426 Goal "!!i::nat. i-j-k = i-k-j"; |
415 Goal "!!i::nat. i-j-k = i-k-j"; |
427 by (simp_tac (simpset() addsimps [diff_diff_left, add_commute]) 1); |
416 by (simp_tac (simpset() addsimps [diff_diff_left, add_commute]) 1); |
428 qed "diff_commute"; |
417 qed "diff_commute"; |
429 |
418 |
478 Goal "i < j ==> ? k. 0<k & i+k = j"; |
467 Goal "i < j ==> ? k. 0<k & i+k = j"; |
479 by (res_inst_tac [("x","j - i")] exI 1); |
468 by (res_inst_tac [("x","j - i")] exI 1); |
480 by (asm_simp_tac (simpset() addsimps [add_diff_inverse, less_not_sym]) 1); |
469 by (asm_simp_tac (simpset() addsimps [add_diff_inverse, less_not_sym]) 1); |
481 qed "less_imp_add_positive"; |
470 qed "less_imp_add_positive"; |
482 |
471 |
483 Goal "Suc(m)-n = (if m<n then 0 else Suc(m-n))"; |
|
484 by (simp_tac (simpset() addsimps [leI, Suc_le_eq, Suc_diff_le]) 1); |
|
485 qed "if_Suc_diff_le"; |
|
486 |
|
487 Goal "Suc(m)-n <= Suc(m-n)"; |
|
488 by (simp_tac (simpset() addsimps [if_Suc_diff_le]) 1); |
|
489 qed "diff_Suc_le_Suc_diff"; |
|
490 |
|
491 Goal "P(k) --> (!n. P(Suc(n))--> P(n)) --> P(k-i)"; |
472 Goal "P(k) --> (!n. P(Suc(n))--> P(n)) --> P(k-i)"; |
492 by (res_inst_tac [("m","k"),("n","i")] diff_induct 1); |
473 by (res_inst_tac [("m","k"),("n","i")] diff_induct 1); |
493 by (ALLGOALS (Clarify_tac THEN' Simp_tac THEN' TRY o Blast_tac)); |
474 by (ALLGOALS (Clarify_tac THEN' Simp_tac THEN' TRY o Blast_tac)); |
494 qed "zero_induct_lemma"; |
475 qed "zero_induct_lemma"; |
495 |
476 |
508 Goal "(m+k) - (n+k) = m - (n::nat)"; |
489 Goal "(m+k) - (n+k) = m - (n::nat)"; |
509 val add_commute_k = read_instantiate [("n","k")] add_commute; |
490 val add_commute_k = read_instantiate [("n","k")] add_commute; |
510 by (asm_simp_tac (simpset() addsimps [add_commute_k]) 1); |
491 by (asm_simp_tac (simpset() addsimps [add_commute_k]) 1); |
511 qed "diff_cancel2"; |
492 qed "diff_cancel2"; |
512 Addsimps [diff_cancel2]; |
493 Addsimps [diff_cancel2]; |
513 |
|
514 (*From Clemens Ballarin, proof by lcp*) |
|
515 Goal "[| k<=n; n<=m |] ==> (m-k) - (n-k) = m-(n::nat)"; |
|
516 by (REPEAT (etac rev_mp 1)); |
|
517 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); |
|
518 by (ALLGOALS Asm_simp_tac); |
|
519 (*a confluence problem*) |
|
520 by (asm_simp_tac (simpset() addsimps [Suc_diff_le, le_Suc_eq]) 1); |
|
521 qed "diff_right_cancel"; |
|
522 |
494 |
523 Goal "n - (n+m) = 0"; |
495 Goal "n - (n+m) = 0"; |
524 by (induct_tac "n" 1); |
496 by (induct_tac "n" 1); |
525 by (ALLGOALS Asm_simp_tac); |
497 by (ALLGOALS Asm_simp_tac); |
526 qed "diff_add_0"; |
498 qed "diff_add_0"; |
1105 Goal "[| 0<n; 0<m |] ==> m - n < m"; |
1077 Goal "[| 0<n; 0<m |] ==> m - n < m"; |
1106 by (arith_tac 1); |
1078 by (arith_tac 1); |
1107 qed "diff_less"; |
1079 qed "diff_less"; |
1108 |
1080 |
1109 |
1081 |
|
1082 (*** Reducting subtraction to addition ***) |
|
1083 |
|
1084 (*Intended for use with linear arithmetic, but useful in its own right*) |
|
1085 Goal "P (x-y) = (ALL z. (x<y --> P 0) & (x = y+z --> P z))"; |
|
1086 by (case_tac "x<y" 1); |
|
1087 by (auto_tac (claset(), simpset() addsimps [diff_is_0_eq RS iffD2])); |
|
1088 qed "split_diff"; |
|
1089 |
|
1090 val remove_diff_ss = |
|
1091 simpset() |
|
1092 delsimps ex_simps@all_simps |
|
1093 addsimps [le_diff_conv2, le_diff_conv, le_imp_diff_is_add, |
|
1094 diff_diff_right] |
|
1095 addcongs [conj_cong] |
|
1096 addsplits [split_diff]; |
|
1097 |
|
1098 Goal "n<=(l::nat) --> Suc l - n + m = Suc (l - n + m)"; |
|
1099 by (simp_tac remove_diff_ss 1); |
|
1100 qed_spec_mp "Suc_diff_add_le"; |
|
1101 |
|
1102 Goal "i<n ==> n - Suc i < n - i"; |
|
1103 by (asm_simp_tac remove_diff_ss 1); |
|
1104 qed "diff_Suc_less_diff"; |
|
1105 |
|
1106 Goal "Suc(m)-n = (if m<n then 0 else Suc(m-n))"; |
|
1107 by (simp_tac remove_diff_ss 1); |
|
1108 qed "if_Suc_diff_le"; |
|
1109 |
|
1110 Goal "Suc(m)-n <= Suc(m-n)"; |
|
1111 by (simp_tac remove_diff_ss 1); |
|
1112 qed "diff_Suc_le_Suc_diff"; |
|
1113 |
|
1114 Goal "[| k<=n; n<=m |] ==> (m-k) - (n-k) = m-(n::nat)"; |
|
1115 by (asm_simp_tac remove_diff_ss 1); |
|
1116 qed "diff_right_cancel"; |
|
1117 |
|
1118 |
1110 (** (Anti)Monotonicity of subtraction -- by Stephan Merz **) |
1119 (** (Anti)Monotonicity of subtraction -- by Stephan Merz **) |
1111 |
1120 |
1112 (* Monotonicity of subtraction in first argument *) |
1121 (* Monotonicity of subtraction in first argument *) |
1113 Goal "m <= (n::nat) ==> (m-l) <= (n-l)"; |
1122 Goal "m <= (n::nat) ==> (m-l) <= (n-l)"; |
1114 by (arith_tac 1); |
1123 by (asm_simp_tac remove_diff_ss 1); |
1115 qed "diff_le_mono"; |
1124 qed "diff_le_mono"; |
1116 |
1125 |
1117 Goal "m <= (n::nat) ==> (l-n) <= (l-m)"; |
1126 Goal "m <= (n::nat) ==> (l-n) <= (l-m)"; |
1118 by (arith_tac 1); |
1127 by (asm_simp_tac remove_diff_ss 1); |
1119 qed "diff_le_mono2"; |
1128 qed "diff_le_mono2"; |
1120 |
|
1121 |
1129 |
1122 (*This proof requires natdiff_cancel_sums*) |
1130 (*This proof requires natdiff_cancel_sums*) |
1123 Goal "[| m < (n::nat); m<l |] ==> (l-n) < (l-m)"; |
1131 Goal "[| m < (n::nat); m<l |] ==> (l-n) < (l-m)"; |
1124 by (arith_tac 1); |
1132 by (asm_simp_tac remove_diff_ss 1); |
1125 qed "diff_less_mono2"; |
1133 qed "diff_less_mono2"; |
1126 |
1134 |
1127 Goal "[| m-n = 0; n-m = 0 |] ==> m=n"; |
1135 Goal "[| m-n = 0; n-m = 0 |] ==> m=n"; |
1128 by (arith_tac 1); |
1136 by (asm_full_simp_tac remove_diff_ss 1); |
1129 qed "diffs0_imp_equal"; |
1137 qed "diffs0_imp_equal"; |