src/HOL/Arith.ML
author wenzelm
Thu, 27 Jul 2000 18:23:12 +0200
changeset 9450 c97dba47e504
parent 9436 62bb04ab4b01
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1398
diff changeset
     1
(*  Title:      HOL/Arith.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1398
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
4736
f7d3b9aec7a1 New laws, mostly generalizing old "pred" ones
paulson
parents: 4732
diff changeset
     4
    Copyright   1998  University of Cambridge
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9073
diff changeset
     6
Further proofs about elementary arithmetic, using the arithmetic proof
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9073
diff changeset
     7
procedures.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9073
diff changeset
    10
(*legacy ...*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9073
diff changeset
    11
structure Arith = struct val thy = the_context () end;
1713
79b4ef7832b5 New cancellation and monotonicity laws, about
paulson
parents: 1660
diff changeset
    12
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5771
diff changeset
    13
8833
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    14
Goal "m <= m*(m::nat)";
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    15
by (induct_tac "m" 1);
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    16
by Auto_tac;
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    17
qed "le_square";
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    18
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    19
Goal "(m::nat) <= m*(m*m)";
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    20
by (induct_tac "m" 1);
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    21
by Auto_tac;
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    22
qed "le_cube";
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    23
d6122f13b8a6 moved le_square, proved le_cube
paulson
parents: 8740
diff changeset
    24
4736
f7d3b9aec7a1 New laws, mostly generalizing old "pred" ones
paulson
parents: 4732
diff changeset
    25
(*** Subtraction laws -- mostly from Clemens Ballarin ***)
3234
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    26
5429
0833486c23ce tidying
paulson
parents: 5427
diff changeset
    27
Goal "[| a < (b::nat); c <= a |] ==> a-c < b-c";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    28
by (arith_tac 1);
3234
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    29
qed "diff_less_mono";
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    30
5427
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5414
diff changeset
    31
Goal "(i < j-k) = (i+k < (j::nat))";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    32
by (arith_tac 1);
5427
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5414
diff changeset
    33
qed "less_diff_conv";
26c9a7c0b36b Arith: less_diff_conv
nipkow
parents: 5414
diff changeset
    34
5497
497215d66441 theorem le_diff_conv2; tidying and expandshort
paulson
parents: 5485
diff changeset
    35
Goal "(j-k <= (i::nat)) = (j <= i+k)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    36
by (arith_tac 1);
5485
0cd451e46a20 new theorem le_diff_conv
paulson
parents: 5429
diff changeset
    37
qed "le_diff_conv";
0cd451e46a20 new theorem le_diff_conv
paulson
parents: 5429
diff changeset
    38
5497
497215d66441 theorem le_diff_conv2; tidying and expandshort
paulson
parents: 5485
diff changeset
    39
Goal "k <= j ==> (i <= j-k) = (i+k <= (j::nat))";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    40
by (arith_tac 1);
5497
497215d66441 theorem le_diff_conv2; tidying and expandshort
paulson
parents: 5485
diff changeset
    41
qed "le_diff_conv2";
497215d66441 theorem le_diff_conv2; tidying and expandshort
paulson
parents: 5485
diff changeset
    42
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5078
diff changeset
    43
Goal "Suc i <= n ==> Suc (n - Suc i) = n - i";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    44
by (arith_tac 1);
3234
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    45
qed "Suc_diff_Suc";
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    46
5429
0833486c23ce tidying
paulson
parents: 5427
diff changeset
    47
Goal "i <= (n::nat) ==> n - (n - i) = i";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    48
by (arith_tac 1);
3234
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    49
qed "diff_diff_cancel";
3381
2bac33ec2b0d New theorems le_add_diff_inverse, le_add_diff_inverse2
paulson
parents: 3366
diff changeset
    50
Addsimps [diff_diff_cancel];
3234
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    51
5429
0833486c23ce tidying
paulson
parents: 5427
diff changeset
    52
Goal "k <= (n::nat) ==> m <= n + m - k";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    53
by (arith_tac 1);
3234
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    54
qed "le_add_diff";
503f4c8c29eb New theorems from Hoare/Arith2.ML
paulson
parents: 2922
diff changeset
    55
5356
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    56
Goal "m-1 < n ==> m <= n";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    57
by (arith_tac 1);
5356
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    58
qed "pred_less_imp_le";
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    59
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    60
Goal "j<=i ==> i - j < Suc i - j";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    61
by (arith_tac 1);
5356
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    62
qed "diff_less_Suc_diff";
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    63
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    64
Goal "i - j <= Suc i - j";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    65
by (arith_tac 1);
5356
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    66
qed "diff_le_Suc_diff";
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    67
AddIffs [diff_le_Suc_diff];
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    68
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    69
Goal "n - Suc i <= n - i";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    70
by (arith_tac 1);
5356
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    71
qed "diff_Suc_le_diff";
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    72
AddIffs [diff_Suc_le_diff];
6ef114ba5b55 new theorems
paulson
parents: 5333
diff changeset
    73
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
    74
Goal "!!m::nat. 0 < n ==> (m <= n-1) = (m<n)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    75
by (arith_tac 1);
5409
e97558ee8e76 Two new subtraction lemmas
paulson
parents: 5356
diff changeset
    76
qed "le_pred_eq";
e97558ee8e76 Two new subtraction lemmas
paulson
parents: 5356
diff changeset
    77
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
    78
Goal "!!m::nat. 0 < n ==> (m-1 < n) = (m<=n)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    79
by (arith_tac 1);
5409
e97558ee8e76 Two new subtraction lemmas
paulson
parents: 5356
diff changeset
    80
qed "less_pred_eq";
e97558ee8e76 Two new subtraction lemmas
paulson
parents: 5356
diff changeset
    81
7059
71e9ea2198e0 a stronger diff_less and no more le_diff_less
paulson
parents: 7007
diff changeset
    82
(*Replaces the previous diff_less and le_diff_less, which had the stronger
71e9ea2198e0 a stronger diff_less and no more le_diff_less
paulson
parents: 7007
diff changeset
    83
  second premise n<=m*)
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
    84
Goal "!!m::nat. [| 0<n; 0<m |] ==> m - n < m";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
    85
by (arith_tac 1);
5414
8a458866637c changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
paulson
parents: 5409
diff changeset
    86
qed "diff_less";
8a458866637c changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
paulson
parents: 5409
diff changeset
    87
8740
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
    88
Goal "j <= (k::nat) ==> (j+i)-k = i-(k-j)";
8859
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
    89
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
8740
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
    90
qed "diff_add_assoc_diff";
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
    91
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
    92
8708
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
    93
(*** Reducing subtraction to addition ***)
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
    94
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
    95
Goal "n<=(l::nat) --> Suc l - n + m = Suc (l - n + m)";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
    96
by (simp_tac (simpset() addsplits [nat_diff_split]) 1);
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
    97
qed_spec_mp "Suc_diff_add_le";
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
    98
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
    99
Goal "i<n ==> n - Suc i < n - i";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   100
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   101
qed "diff_Suc_less_diff";
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   102
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   103
Goal "Suc(m)-n = (if m<n then 0 else Suc(m-n))";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   104
by (simp_tac (simpset() addsplits [nat_diff_split]) 1);
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   105
qed "if_Suc_diff_le";
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   106
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   107
Goal "Suc(m)-n <= Suc(m-n)";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   108
by (simp_tac (simpset() addsplits [nat_diff_split]) 1);
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   109
qed "diff_Suc_le_Suc_diff";
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   110
8740
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
   111
(** Simplification of relational expressions involving subtraction **)
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   112
8740
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
   113
Goal "[| k <= m;  k <= (n::nat) |] ==> ((m-k) - (n-k)) = (m-n)";
8859
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   114
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
8740
fa6c4e4182d9 replaced obsolete diff_right_cancel by diff_diff_eq
paulson
parents: 8708
diff changeset
   115
qed "diff_diff_eq";
8708
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   116
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   117
Goal "[| k <= m;  k <= (n::nat) |] ==> (m-k = n-k) = (m=n)";
8859
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   118
by (auto_tac (claset(), simpset() addsplits [nat_diff_split]));
8708
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   119
qed "eq_diff_iff";
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   120
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   121
Goal "[| k <= m;  k <= (n::nat) |] ==> (m-k < n-k) = (m<n)";
8859
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   122
by (auto_tac (claset(), simpset() addsplits [nat_diff_split]));
8708
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   123
qed "less_diff_iff";
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   124
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   125
Goal "[| k <= m;  k <= (n::nat) |] ==> (m-k <= n-k) = (m<=n)";
8859
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   126
by (auto_tac (claset(), simpset() addsplits [nat_diff_split]));
8708
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   127
qed "le_diff_iff";
2f2d2b4215d6 added some new iff-lemmas; removed some obsolete thms
paulson
parents: 8571
diff changeset
   128
7128
468b6c8b8dc4 split_diff and remove_diff_ss
paulson
parents: 7108
diff changeset
   129
7108
0229ce6735f6 fixed comment;
wenzelm
parents: 7059
diff changeset
   130
(** (Anti)Monotonicity of subtraction -- by Stephan Merz **)
3484
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   131
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   132
(* Monotonicity of subtraction in first argument *)
6055
fdf4638bf726 Version 1.0 of linear nat arithmetic.
nipkow
parents: 5983
diff changeset
   133
Goal "m <= (n::nat) ==> (m-l) <= (n-l)";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   134
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
6055
fdf4638bf726 Version 1.0 of linear nat arithmetic.
nipkow
parents: 5983
diff changeset
   135
qed "diff_le_mono";
3484
1e93eb09ebb9 Added the following lemmas tp Divides and a few others to Arith and NatDef:
nipkow
parents: 3457
diff changeset
   136
5429
0833486c23ce tidying
paulson
parents: 5427
diff changeset
   137
Goal "m <= (n::nat) ==> (l-n) <= (l-m)";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   138
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
6055
fdf4638bf726 Version 1.0 of linear nat arithmetic.
nipkow
parents: 5983
diff changeset
   139
qed "diff_le_mono2";
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5771
diff changeset
   140
6055
fdf4638bf726 Version 1.0 of linear nat arithmetic.
nipkow
parents: 5983
diff changeset
   141
Goal "[| m < (n::nat); m<l |] ==> (l-n) < (l-m)";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   142
by (asm_simp_tac (simpset() addsplits [nat_diff_split]) 1);
6055
fdf4638bf726 Version 1.0 of linear nat arithmetic.
nipkow
parents: 5983
diff changeset
   143
qed "diff_less_mono2";
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5771
diff changeset
   144
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
   145
Goal "!!m::nat. [| m-n = 0; n-m = 0 |] ==>  m=n";
7945
3aca6352f063 got rid of split_diff, which duplicated nat_diff_split, and
paulson
parents: 7622
diff changeset
   146
by (asm_full_simp_tac (simpset() addsplits [nat_diff_split]) 1);
6055
fdf4638bf726 Version 1.0 of linear nat arithmetic.
nipkow
parents: 5983
diff changeset
   147
qed "diffs0_imp_equal";
8352
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   148
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   149
(** Lemmas for ex/Factorization **)
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   150
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
   151
Goal "!!m::nat. [| 1<n; 1<m |] ==> 1<m*n";
8442
96023903c2df case_tac now subsumes both boolean and datatype cases;
wenzelm
parents: 8423
diff changeset
   152
by (case_tac "m" 1);
8352
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   153
by Auto_tac;
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   154
qed "one_less_mult"; 
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   155
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
   156
Goal "!!m::nat. [| 1<n; 1<m |] ==> n<m*n";
8442
96023903c2df case_tac now subsumes both boolean and datatype cases;
wenzelm
parents: 8423
diff changeset
   157
by (case_tac "m" 1);
8352
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   158
by Auto_tac;
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   159
qed "n_less_m_mult_n"; 
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   160
8935
548901d05a0e added type constraint ::nat because 0 is now overloaded
paulson
parents: 8859
diff changeset
   161
Goal "!!m::nat. [| 1<n; 1<m |] ==> n<n*m";
8442
96023903c2df case_tac now subsumes both boolean and datatype cases;
wenzelm
parents: 8423
diff changeset
   162
by (case_tac "m" 1);
8352
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   163
by Auto_tac;
0fda5ba36934 new lemmas
paulson
parents: 8252
diff changeset
   164
qed "n_less_n_mult_m"; 
8859
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   165
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   166
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   167
(** Rewriting to pull differences out **)
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   168
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   169
Goal "k<=j --> i - (j - k) = i + (k::nat) - j";
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   170
by (arith_tac 1);
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   171
qed "diff_diff_right";
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   172
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   173
Goal "k <= j ==> m - Suc (j - k) = m + k - Suc j";
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   174
by (arith_tac 1);
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   175
qed "diff_Suc_diff_eq1"; 
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   176
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   177
Goal "k <= j ==> Suc (j - k) - m = Suc j - (k + m)";
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   178
by (arith_tac 1);
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   179
qed "diff_Suc_diff_eq2"; 
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   180
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   181
(*The others are
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   182
      i - j - k = i - (j + k),
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   183
      k <= j ==> j - k + i = j + i - k,
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   184
      k <= j ==> i + (j - k) = i + j - k *)
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   185
Addsimps [diff_diff_left, diff_diff_right, diff_add_assoc2 RS sym, 
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   186
	  diff_add_assoc RS sym, diff_Suc_diff_eq1, diff_Suc_diff_eq2];
b1ea21d70c85 deleted some redundant simprules
paulson
parents: 8833
diff changeset
   187