| author | wenzelm | 
| Mon, 08 Feb 1999 17:30:22 +0100 | |
| changeset 6260 | a8010d459ef7 | 
| parent 6157 | 29942d3a1818 | 
| child 6301 | 08245f5a436d | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/Arith.ML | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 4736 | 4 | Copyright 1998 University of Cambridge | 
| 923 | 5 | |
| 6 | Proofs about elementary arithmetic: addition, multiplication, etc. | |
| 3234 | 7 | Some from the Hoare example from Norbert Galm | 
| 923 | 8 | *) | 
| 9 | ||
| 10 | (*** Basic rewrite rules for the arithmetic operators ***) | |
| 11 | ||
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3842diff
changeset | 12 | |
| 923 | 13 | (** Difference **) | 
| 14 | ||
| 4732 | 15 | qed_goal "diff_0_eq_0" thy | 
| 923 | 16 | "0 - n = 0" | 
| 3339 | 17 | (fn _ => [induct_tac "n" 1, ALLGOALS Asm_simp_tac]); | 
| 923 | 18 | |
| 5429 | 19 | (*Must simplify BEFORE the induction! (Else we get a critical pair) | 
| 923 | 20 | Suc(m) - Suc(n) rewrites to pred(Suc(m) - n) *) | 
| 4732 | 21 | qed_goal "diff_Suc_Suc" thy | 
| 923 | 22 | "Suc(m) - Suc(n) = m - n" | 
| 23 | (fn _ => | |
| 3339 | 24 | [Simp_tac 1, induct_tac "n" 1, ALLGOALS Asm_simp_tac]); | 
| 923 | 25 | |
| 2682 
13cdbf95ed92
minor changes due to new primrec definitions for +,-,*
 pusch parents: 
2498diff
changeset | 26 | Addsimps [diff_0_eq_0, diff_Suc_Suc]; | 
| 923 | 27 | |
| 4360 | 28 | (* Could be (and is, below) generalized in various ways; | 
| 29 | However, none of the generalizations are currently in the simpset, | |
| 30 | and I dread to think what happens if I put them in *) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 31 | Goal "0 < n ==> Suc(n-1) = n"; | 
| 5183 | 32 | by (asm_simp_tac (simpset() addsplits [nat.split]) 1); | 
| 4360 | 33 | qed "Suc_pred"; | 
| 34 | Addsimps [Suc_pred]; | |
| 35 | ||
| 36 | Delsimps [diff_Suc]; | |
| 37 | ||
| 923 | 38 | |
| 39 | (**** Inductive properties of the operators ****) | |
| 40 | ||
| 41 | (*** Addition ***) | |
| 42 | ||
| 4732 | 43 | qed_goal "add_0_right" thy "m + 0 = m" | 
| 3339 | 44 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 923 | 45 | |
| 4732 | 46 | qed_goal "add_Suc_right" thy "m + Suc(n) = Suc(m+n)" | 
| 3339 | 47 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 923 | 48 | |
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 49 | Addsimps [add_0_right,add_Suc_right]; | 
| 923 | 50 | |
| 51 | (*Associative law for addition*) | |
| 4732 | 52 | qed_goal "add_assoc" thy "(m + n) + k = m + ((n + k)::nat)" | 
| 3339 | 53 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 923 | 54 | |
| 55 | (*Commutative law for addition*) | |
| 4732 | 56 | qed_goal "add_commute" thy "m + n = n + (m::nat)" | 
| 3339 | 57 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 923 | 58 | |
| 4732 | 59 | qed_goal "add_left_commute" thy "x+(y+z)=y+((x+z)::nat)" | 
| 923 | 60 | (fn _ => [rtac (add_commute RS trans) 1, rtac (add_assoc RS trans) 1, | 
| 61 | rtac (add_commute RS arg_cong) 1]); | |
| 62 | ||
| 63 | (*Addition is an AC-operator*) | |
| 64 | val add_ac = [add_assoc, add_commute, add_left_commute]; | |
| 65 | ||
| 5429 | 66 | Goal "(k + m = k + n) = (m=(n::nat))"; | 
| 3339 | 67 | by (induct_tac "k" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 68 | by (Simp_tac 1); | 
| 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 69 | by (Asm_simp_tac 1); | 
| 923 | 70 | qed "add_left_cancel"; | 
| 71 | ||
| 5429 | 72 | Goal "(m + k = n + k) = (m=(n::nat))"; | 
| 3339 | 73 | by (induct_tac "k" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 74 | by (Simp_tac 1); | 
| 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 75 | by (Asm_simp_tac 1); | 
| 923 | 76 | qed "add_right_cancel"; | 
| 77 | ||
| 5429 | 78 | Goal "(k + m <= k + n) = (m<=(n::nat))"; | 
| 3339 | 79 | by (induct_tac "k" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 80 | by (Simp_tac 1); | 
| 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 81 | by (Asm_simp_tac 1); | 
| 923 | 82 | qed "add_left_cancel_le"; | 
| 83 | ||
| 5429 | 84 | Goal "(k + m < k + n) = (m<(n::nat))"; | 
| 3339 | 85 | by (induct_tac "k" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 86 | by (Simp_tac 1); | 
| 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 87 | by (Asm_simp_tac 1); | 
| 923 | 88 | qed "add_left_cancel_less"; | 
| 89 | ||
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 90 | Addsimps [add_left_cancel, add_right_cancel, | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 91 | add_left_cancel_le, add_left_cancel_less]; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 92 | |
| 3339 | 93 | (** Reasoning about m+0=0, etc. **) | 
| 94 | ||
| 5069 | 95 | Goal "(m+n = 0) = (m=0 & n=0)"; | 
| 5598 | 96 | by (exhaust_tac "m" 1); | 
| 97 | by (Auto_tac); | |
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 98 | qed "add_is_0"; | 
| 4360 | 99 | AddIffs [add_is_0]; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 100 | |
| 5598 | 101 | Goal "(0 = m+n) = (m=0 & n=0)"; | 
| 102 | by (exhaust_tac "m" 1); | |
| 103 | by (Auto_tac); | |
| 104 | qed "zero_is_add"; | |
| 105 | AddIffs [zero_is_add]; | |
| 106 | ||
| 107 | Goal "(m+n=1) = (m=1 & n=0 | m=0 & n=1)"; | |
| 108 | by(exhaust_tac "m" 1); | |
| 109 | by(Auto_tac); | |
| 110 | qed "add_is_1"; | |
| 111 | ||
| 112 | Goal "(1=m+n) = (m=1 & n=0 | m=0 & n=1)"; | |
| 113 | by(exhaust_tac "m" 1); | |
| 114 | by(Auto_tac); | |
| 115 | qed "one_is_add"; | |
| 116 | ||
| 5069 | 117 | Goal "(0<m+n) = (0<m | 0<n)"; | 
| 4423 | 118 | by (simp_tac (simpset() delsimps [neq0_conv] addsimps [neq0_conv RS sym]) 1); | 
| 4360 | 119 | qed "add_gr_0"; | 
| 120 | AddIffs [add_gr_0]; | |
| 121 | ||
| 122 | (* FIXME: really needed?? *) | |
| 5069 | 123 | Goal "((m+n)-1 = 0) = (m=0 & n-1 = 0 | m-1 = 0 & n=0)"; | 
| 4360 | 124 | by (exhaust_tac "m" 1); | 
| 4089 | 125 | by (ALLGOALS (fast_tac (claset() addss (simpset())))); | 
| 3293 | 126 | qed "pred_add_is_0"; | 
| 6075 | 127 | (*Addsimps [pred_add_is_0];*) | 
| 3293 | 128 | |
| 5429 | 129 | (* Could be generalized, eg to "k<n ==> m+(n-(Suc k)) = (m+n)-(Suc k)" *) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 130 | Goal "0<n ==> m + (n-1) = (m+n)-1"; | 
| 4360 | 131 | by (exhaust_tac "m" 1); | 
| 6075 | 132 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_Suc, Suc_n_not_n] | 
| 5183 | 133 | addsplits [nat.split]))); | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 134 | qed "add_pred"; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 135 | Addsimps [add_pred]; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 136 | |
| 5429 | 137 | Goal "m + n = m ==> n = 0"; | 
| 5078 | 138 | by (dtac (add_0_right RS ssubst) 1); | 
| 139 | by (asm_full_simp_tac (simpset() addsimps [add_assoc] | |
| 140 | delsimps [add_0_right]) 1); | |
| 141 | qed "add_eq_self_zero"; | |
| 142 | ||
| 1626 | 143 | |
| 923 | 144 | (**** Additional theorems about "less than" ****) | 
| 145 | ||
| 5078 | 146 | (*Deleted less_natE; instead use less_eq_Suc_add RS exE*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 147 | Goal "m<n --> (? k. n=Suc(m+k))"; | 
| 3339 | 148 | by (induct_tac "n" 1); | 
| 5604 | 149 | by (ALLGOALS (simp_tac (simpset() addsimps [order_le_less]))); | 
| 4089 | 150 | by (blast_tac (claset() addSEs [less_SucE] | 
| 5497 | 151 | addSIs [add_0_right RS sym, add_Suc_right RS sym]) 1); | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1475diff
changeset | 152 | qed_spec_mp "less_eq_Suc_add"; | 
| 923 | 153 | |
| 5069 | 154 | Goal "n <= ((m + n)::nat)"; | 
| 3339 | 155 | by (induct_tac "m" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 156 | by (ALLGOALS Simp_tac); | 
| 5983 | 157 | by (etac le_SucI 1); | 
| 923 | 158 | qed "le_add2"; | 
| 159 | ||
| 5069 | 160 | Goal "n <= ((n + m)::nat)"; | 
| 4089 | 161 | by (simp_tac (simpset() addsimps add_ac) 1); | 
| 923 | 162 | by (rtac le_add2 1); | 
| 163 | qed "le_add1"; | |
| 164 | ||
| 165 | bind_thm ("less_add_Suc1", (lessI RS (le_add1 RS le_less_trans)));
 | |
| 166 | bind_thm ("less_add_Suc2", (lessI RS (le_add2 RS le_less_trans)));
 | |
| 167 | ||
| 5429 | 168 | Goal "(m<n) = (? k. n=Suc(m+k))"; | 
| 169 | by (blast_tac (claset() addSIs [less_add_Suc1, less_eq_Suc_add]) 1); | |
| 170 | qed "less_iff_Suc_add"; | |
| 171 | ||
| 172 | ||
| 923 | 173 | (*"i <= j ==> i <= j+m"*) | 
| 174 | bind_thm ("trans_le_add1", le_add1 RSN (2,le_trans));
 | |
| 175 | ||
| 176 | (*"i <= j ==> i <= m+j"*) | |
| 177 | bind_thm ("trans_le_add2", le_add2 RSN (2,le_trans));
 | |
| 178 | ||
| 179 | (*"i < j ==> i < j+m"*) | |
| 180 | bind_thm ("trans_less_add1", le_add1 RSN (2,less_le_trans));
 | |
| 181 | ||
| 182 | (*"i < j ==> i < m+j"*) | |
| 183 | bind_thm ("trans_less_add2", le_add2 RSN (2,less_le_trans));
 | |
| 184 | ||
| 5654 | 185 | Goal "i+j < (k::nat) --> i<k"; | 
| 3339 | 186 | by (induct_tac "j" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 187 | by (ALLGOALS Asm_simp_tac); | 
| 5983 | 188 | by(blast_tac (claset() addDs [Suc_lessD]) 1); | 
| 5654 | 189 | qed_spec_mp "add_lessD1"; | 
| 1152 | 190 | |
| 5429 | 191 | Goal "~ (i+j < (i::nat))"; | 
| 3457 | 192 | by (rtac notI 1); | 
| 193 | by (etac (add_lessD1 RS less_irrefl) 1); | |
| 3234 | 194 | qed "not_add_less1"; | 
| 195 | ||
| 5429 | 196 | Goal "~ (j+i < (i::nat))"; | 
| 4089 | 197 | by (simp_tac (simpset() addsimps [add_commute, not_add_less1]) 1); | 
| 3234 | 198 | qed "not_add_less2"; | 
| 199 | AddIffs [not_add_less1, not_add_less2]; | |
| 200 | ||
| 5069 | 201 | Goal "m+k<=n --> m<=(n::nat)"; | 
| 3339 | 202 | by (induct_tac "k" 1); | 
| 5497 | 203 | by (ALLGOALS (asm_simp_tac (simpset() addsimps le_simps))); | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1475diff
changeset | 204 | qed_spec_mp "add_leD1"; | 
| 923 | 205 | |
| 5429 | 206 | Goal "m+k<=n ==> k<=(n::nat)"; | 
| 4089 | 207 | by (full_simp_tac (simpset() addsimps [add_commute]) 1); | 
| 2498 | 208 | by (etac add_leD1 1); | 
| 209 | qed_spec_mp "add_leD2"; | |
| 210 | ||
| 5429 | 211 | Goal "m+k<=n ==> m<=n & k<=(n::nat)"; | 
| 4089 | 212 | by (blast_tac (claset() addDs [add_leD1, add_leD2]) 1); | 
| 2498 | 213 | bind_thm ("add_leE", result() RS conjE);
 | 
| 214 | ||
| 5429 | 215 | (*needs !!k for add_ac to work*) | 
| 216 | Goal "!!k:: nat. [| k<l; m+l = k+n |] ==> m<n"; | |
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5654diff
changeset | 217 | by (force_tac (claset(), | 
| 5497 | 218 | simpset() delsimps [add_Suc_right] | 
| 5537 | 219 | addsimps [less_iff_Suc_add, | 
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5654diff
changeset | 220 | add_Suc_right RS sym] @ add_ac) 1); | 
| 923 | 221 | qed "less_add_eq_less"; | 
| 222 | ||
| 223 | ||
| 1713 | 224 | (*** Monotonicity of Addition ***) | 
| 923 | 225 | |
| 226 | (*strict, in 1st argument*) | |
| 5429 | 227 | Goal "i < j ==> i + k < j + (k::nat)"; | 
| 3339 | 228 | by (induct_tac "k" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 229 | by (ALLGOALS Asm_simp_tac); | 
| 923 | 230 | qed "add_less_mono1"; | 
| 231 | ||
| 232 | (*strict, in both arguments*) | |
| 5429 | 233 | Goal "[|i < j; k < l|] ==> i + k < j + (l::nat)"; | 
| 923 | 234 | by (rtac (add_less_mono1 RS less_trans) 1); | 
| 1198 | 235 | by (REPEAT (assume_tac 1)); | 
| 3339 | 236 | by (induct_tac "j" 1); | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1198diff
changeset | 237 | by (ALLGOALS Asm_simp_tac); | 
| 923 | 238 | qed "add_less_mono"; | 
| 239 | ||
| 240 | (*A [clumsy] way of lifting < monotonicity to <= monotonicity *) | |
| 5316 | 241 | val [lt_mono,le] = Goal | 
| 1465 | 242 | "[| !!i j::nat. i<j ==> f(i) < f(j); \ | 
| 243 | \ i <= j \ | |
| 923 | 244 | \ |] ==> f(i) <= (f(j)::nat)"; | 
| 245 | by (cut_facts_tac [le] 1); | |
| 5604 | 246 | by (asm_full_simp_tac (simpset() addsimps [order_le_less]) 1); | 
| 4089 | 247 | by (blast_tac (claset() addSIs [lt_mono]) 1); | 
| 923 | 248 | qed "less_mono_imp_le_mono"; | 
| 249 | ||
| 250 | (*non-strict, in 1st argument*) | |
| 5429 | 251 | Goal "i<=j ==> i + k <= j + (k::nat)"; | 
| 3842 | 252 | by (res_inst_tac [("f", "%j. j+k")] less_mono_imp_le_mono 1);
 | 
| 1552 | 253 | by (etac add_less_mono1 1); | 
| 923 | 254 | by (assume_tac 1); | 
| 255 | qed "add_le_mono1"; | |
| 256 | ||
| 257 | (*non-strict, in both arguments*) | |
| 5429 | 258 | Goal "[|i<=j; k<=l |] ==> i + k <= j + (l::nat)"; | 
| 923 | 259 | by (etac (add_le_mono1 RS le_trans) 1); | 
| 4089 | 260 | by (simp_tac (simpset() addsimps [add_commute]) 1); | 
| 923 | 261 | qed "add_le_mono"; | 
| 1713 | 262 | |
| 3234 | 263 | |
| 264 | (*** Multiplication ***) | |
| 265 | ||
| 266 | (*right annihilation in product*) | |
| 4732 | 267 | qed_goal "mult_0_right" thy "m * 0 = 0" | 
| 3339 | 268 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 3234 | 269 | |
| 3293 | 270 | (*right successor law for multiplication*) | 
| 4732 | 271 | qed_goal "mult_Suc_right" thy "m * Suc(n) = m + (m * n)" | 
| 3339 | 272 | (fn _ => [induct_tac "m" 1, | 
| 4089 | 273 | ALLGOALS(asm_simp_tac (simpset() addsimps add_ac))]); | 
| 3234 | 274 | |
| 3293 | 275 | Addsimps [mult_0_right, mult_Suc_right]; | 
| 3234 | 276 | |
| 5069 | 277 | Goal "1 * n = n"; | 
| 3234 | 278 | by (Asm_simp_tac 1); | 
| 279 | qed "mult_1"; | |
| 280 | ||
| 5069 | 281 | Goal "n * 1 = n"; | 
| 3234 | 282 | by (Asm_simp_tac 1); | 
| 283 | qed "mult_1_right"; | |
| 284 | ||
| 285 | (*Commutative law for multiplication*) | |
| 4732 | 286 | qed_goal "mult_commute" thy "m * n = n * (m::nat)" | 
| 3339 | 287 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 3234 | 288 | |
| 289 | (*addition distributes over multiplication*) | |
| 4732 | 290 | qed_goal "add_mult_distrib" thy "(m + n)*k = (m*k) + ((n*k)::nat)" | 
| 3339 | 291 | (fn _ => [induct_tac "m" 1, | 
| 4089 | 292 | ALLGOALS(asm_simp_tac (simpset() addsimps add_ac))]); | 
| 3234 | 293 | |
| 4732 | 294 | qed_goal "add_mult_distrib2" thy "k*(m + n) = (k*m) + ((k*n)::nat)" | 
| 3339 | 295 | (fn _ => [induct_tac "m" 1, | 
| 4089 | 296 | ALLGOALS(asm_simp_tac (simpset() addsimps add_ac))]); | 
| 3234 | 297 | |
| 298 | (*Associative law for multiplication*) | |
| 4732 | 299 | qed_goal "mult_assoc" thy "(m * n) * k = m * ((n * k)::nat)" | 
| 3339 | 300 | (fn _ => [induct_tac "m" 1, | 
| 4089 | 301 | ALLGOALS (asm_simp_tac (simpset() addsimps [add_mult_distrib]))]); | 
| 3234 | 302 | |
| 4732 | 303 | qed_goal "mult_left_commute" thy "x*(y*z) = y*((x*z)::nat)" | 
| 3234 | 304 | (fn _ => [rtac trans 1, rtac mult_commute 1, rtac trans 1, | 
| 305 | rtac mult_assoc 1, rtac (mult_commute RS arg_cong) 1]); | |
| 306 | ||
| 307 | val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; | |
| 308 | ||
| 5069 | 309 | Goal "(m*n = 0) = (m=0 | n=0)"; | 
| 3339 | 310 | by (induct_tac "m" 1); | 
| 311 | by (induct_tac "n" 2); | |
| 3293 | 312 | by (ALLGOALS Asm_simp_tac); | 
| 313 | qed "mult_is_0"; | |
| 314 | Addsimps [mult_is_0]; | |
| 315 | ||
| 5429 | 316 | Goal "m <= m*(m::nat)"; | 
| 4158 | 317 | by (induct_tac "m" 1); | 
| 318 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_assoc RS sym]))); | |
| 319 | by (etac (le_add2 RSN (2,le_trans)) 1); | |
| 320 | qed "le_square"; | |
| 321 | ||
| 3234 | 322 | |
| 323 | (*** Difference ***) | |
| 324 | ||
| 325 | ||
| 4732 | 326 | qed_goal "diff_self_eq_0" thy "m - m = 0" | 
| 3339 | 327 | (fn _ => [induct_tac "m" 1, ALLGOALS Asm_simp_tac]); | 
| 3234 | 328 | Addsimps [diff_self_eq_0]; | 
| 329 | ||
| 330 | (*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *) | |
| 5069 | 331 | Goal "~ m<n --> n+(m-n) = (m::nat)"; | 
| 3234 | 332 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 3352 | 333 | by (ALLGOALS Asm_simp_tac); | 
| 3381 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 334 | qed_spec_mp "add_diff_inverse"; | 
| 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 335 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 336 | Goal "n<=m ==> n+(m-n) = (m::nat)"; | 
| 4089 | 337 | by (asm_simp_tac (simpset() addsimps [add_diff_inverse, not_less_iff_le]) 1); | 
| 3381 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 338 | qed "le_add_diff_inverse"; | 
| 3234 | 339 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 340 | Goal "n<=m ==> (m-n)+n = (m::nat)"; | 
| 4089 | 341 | by (asm_simp_tac (simpset() addsimps [le_add_diff_inverse, add_commute]) 1); | 
| 3381 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 342 | qed "le_add_diff_inverse2"; | 
| 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 343 | |
| 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 344 | Addsimps [le_add_diff_inverse, le_add_diff_inverse2]; | 
| 3234 | 345 | |
| 346 | ||
| 347 | (*** More results about difference ***) | |
| 348 | ||
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 349 | Goal "n <= m ==> Suc(m)-n = Suc(m-n)"; | 
| 5316 | 350 | by (etac rev_mp 1); | 
| 3352 | 351 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 352 | by (ALLGOALS Asm_simp_tac); | |
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 353 | qed "Suc_diff_le"; | 
| 3352 | 354 | |
| 5429 | 355 | Goal "n<=(l::nat) --> Suc l - n + m = Suc (l - n + m)"; | 
| 356 | by (res_inst_tac [("m","n"),("n","l")] diff_induct 1);
 | |
| 357 | by (ALLGOALS Asm_simp_tac); | |
| 358 | qed_spec_mp "Suc_diff_add_le"; | |
| 359 | ||
| 5069 | 360 | Goal "m - n < Suc(m)"; | 
| 3234 | 361 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 362 | by (etac less_SucE 3); | |
| 4089 | 363 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq]))); | 
| 3234 | 364 | qed "diff_less_Suc"; | 
| 365 | ||
| 5429 | 366 | Goal "m - n <= (m::nat)"; | 
| 3234 | 367 | by (res_inst_tac [("m","m"), ("n","n")] diff_induct 1);
 | 
| 6075 | 368 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [le_SucI]))); | 
| 3234 | 369 | qed "diff_le_self"; | 
| 3903 
1b29151a1009
New simprule diff_le_self, requiring a new proof of diff_diff_cancel
 paulson parents: 
3896diff
changeset | 370 | Addsimps [diff_le_self]; | 
| 3234 | 371 | |
| 4732 | 372 | (* j<k ==> j-n < k *) | 
| 373 | bind_thm ("less_imp_diff_less", diff_le_self RS le_less_trans);
 | |
| 374 | ||
| 5069 | 375 | Goal "!!i::nat. i-j-k = i - (j+k)"; | 
| 3352 | 376 | by (res_inst_tac [("m","i"),("n","j")] diff_induct 1);
 | 
| 377 | by (ALLGOALS Asm_simp_tac); | |
| 378 | qed "diff_diff_left"; | |
| 379 | ||
| 5069 | 380 | Goal "(Suc m - n) - Suc k = m - n - k"; | 
| 4423 | 381 | by (simp_tac (simpset() addsimps [diff_diff_left]) 1); | 
| 4736 | 382 | qed "Suc_diff_diff"; | 
| 383 | Addsimps [Suc_diff_diff]; | |
| 4360 | 384 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 385 | Goal "0<n ==> n - Suc i < n"; | 
| 5183 | 386 | by (exhaust_tac "n" 1); | 
| 4732 | 387 | by Safe_tac; | 
| 5497 | 388 | by (asm_simp_tac (simpset() addsimps le_simps) 1); | 
| 4732 | 389 | qed "diff_Suc_less"; | 
| 390 | Addsimps [diff_Suc_less]; | |
| 391 | ||
| 5329 | 392 | Goal "i<n ==> n - Suc i < n - i"; | 
| 393 | by (exhaust_tac "n" 1); | |
| 5497 | 394 | by (auto_tac (claset(), | 
| 5537 | 395 | simpset() addsimps [Suc_diff_le]@le_simps)); | 
| 5329 | 396 | qed "diff_Suc_less_diff"; | 
| 397 | ||
| 3396 | 398 | (*This and the next few suggested by Florian Kammueller*) | 
| 5069 | 399 | Goal "!!i::nat. i-j-k = i-k-j"; | 
| 4089 | 400 | by (simp_tac (simpset() addsimps [diff_diff_left, add_commute]) 1); | 
| 3352 | 401 | qed "diff_commute"; | 
| 402 | ||
| 5429 | 403 | Goal "k<=j --> j<=i --> i - (j - k) = i - j + (k::nat)"; | 
| 3352 | 404 | by (res_inst_tac [("m","i"),("n","j")] diff_induct 1);
 | 
| 405 | by (ALLGOALS Asm_simp_tac); | |
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 406 | by (asm_simp_tac (simpset() addsimps [Suc_diff_le, le_Suc_eq]) 1); | 
| 3352 | 407 | qed_spec_mp "diff_diff_right"; | 
| 408 | ||
| 5429 | 409 | Goal "k <= (j::nat) --> (i + j) - k = i + (j - k)"; | 
| 3352 | 410 | by (res_inst_tac [("m","j"),("n","k")] diff_induct 1);
 | 
| 411 | by (ALLGOALS Asm_simp_tac); | |
| 412 | qed_spec_mp "diff_add_assoc"; | |
| 413 | ||
| 5429 | 414 | Goal "k <= (j::nat) --> (j + i) - k = i + (j - k)"; | 
| 4732 | 415 | by (asm_simp_tac (simpset() addsimps [add_commute, diff_add_assoc]) 1); | 
| 416 | qed_spec_mp "diff_add_assoc2"; | |
| 417 | ||
| 5429 | 418 | Goal "(n+m) - n = (m::nat)"; | 
| 3339 | 419 | by (induct_tac "n" 1); | 
| 3234 | 420 | by (ALLGOALS Asm_simp_tac); | 
| 421 | qed "diff_add_inverse"; | |
| 422 | Addsimps [diff_add_inverse]; | |
| 423 | ||
| 5429 | 424 | Goal "(m+n) - n = (m::nat)"; | 
| 4089 | 425 | by (simp_tac (simpset() addsimps [diff_add_assoc]) 1); | 
| 3234 | 426 | qed "diff_add_inverse2"; | 
| 427 | Addsimps [diff_add_inverse2]; | |
| 428 | ||
| 5429 | 429 | Goal "i <= (j::nat) ==> (j-i=k) = (j=k+i)"; | 
| 3724 | 430 | by Safe_tac; | 
| 3381 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 431 | by (ALLGOALS Asm_simp_tac); | 
| 3366 | 432 | qed "le_imp_diff_is_add"; | 
| 433 | ||
| 5356 | 434 | Goal "(m-n = 0) = (m <= n)"; | 
| 3234 | 435 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 5497 | 436 | by (ALLGOALS Asm_simp_tac); | 
| 5356 | 437 | qed "diff_is_0_eq"; | 
| 438 | Addsimps [diff_is_0_eq RS iffD2]; | |
| 3234 | 439 | |
| 5333 
ea33e66dcedd
Some new theorems.  zero_less_diff replaces less_imp_diff_positive
 paulson parents: 
5329diff
changeset | 440 | Goal "(0<n-m) = (m<n)"; | 
| 3234 | 441 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 3352 | 442 | by (ALLGOALS Asm_simp_tac); | 
| 5333 
ea33e66dcedd
Some new theorems.  zero_less_diff replaces less_imp_diff_positive
 paulson parents: 
5329diff
changeset | 443 | qed "zero_less_diff"; | 
| 
ea33e66dcedd
Some new theorems.  zero_less_diff replaces less_imp_diff_positive
 paulson parents: 
5329diff
changeset | 444 | Addsimps [zero_less_diff]; | 
| 3234 | 445 | |
| 5333 
ea33e66dcedd
Some new theorems.  zero_less_diff replaces less_imp_diff_positive
 paulson parents: 
5329diff
changeset | 446 | Goal "i < j ==> ? k. 0<k & i+k = j"; | 
| 5078 | 447 | by (res_inst_tac [("x","j - i")] exI 1);
 | 
| 5333 
ea33e66dcedd
Some new theorems.  zero_less_diff replaces less_imp_diff_positive
 paulson parents: 
5329diff
changeset | 448 | by (asm_simp_tac (simpset() addsimps [add_diff_inverse, less_not_sym]) 1); | 
| 5078 | 449 | qed "less_imp_add_positive"; | 
| 450 | ||
| 5069 | 451 | Goal "Suc(m)-n = (if m<n then 0 else Suc(m-n))"; | 
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 452 | by (simp_tac (simpset() addsimps [leI, Suc_le_eq, Suc_diff_le]) 1); | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 453 | qed "if_Suc_diff_le"; | 
| 3234 | 454 | |
| 5069 | 455 | Goal "Suc(m)-n <= Suc(m-n)"; | 
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 456 | by (simp_tac (simpset() addsimps [if_Suc_diff_le]) 1); | 
| 4672 
9d55bc687e1e
New theorem diff_Suc_le_Suc_diff; tidied another proof
 paulson parents: 
4423diff
changeset | 457 | qed "diff_Suc_le_Suc_diff"; | 
| 
9d55bc687e1e
New theorem diff_Suc_le_Suc_diff; tidied another proof
 paulson parents: 
4423diff
changeset | 458 | |
| 5069 | 459 | Goal "P(k) --> (!n. P(Suc(n))--> P(n)) --> P(k-i)"; | 
| 3234 | 460 | by (res_inst_tac [("m","k"),("n","i")] diff_induct 1);
 | 
| 3718 | 461 | by (ALLGOALS (Clarify_tac THEN' Simp_tac THEN' TRY o Blast_tac)); | 
| 3234 | 462 | qed "zero_induct_lemma"; | 
| 463 | ||
| 5316 | 464 | val prems = Goal "[| P(k); !!n. P(Suc(n)) ==> P(n) |] ==> P(0)"; | 
| 3234 | 465 | by (rtac (diff_self_eq_0 RS subst) 1); | 
| 466 | by (rtac (zero_induct_lemma RS mp RS mp) 1); | |
| 467 | by (REPEAT (ares_tac ([impI,allI]@prems) 1)); | |
| 468 | qed "zero_induct"; | |
| 469 | ||
| 5429 | 470 | Goal "(k+m) - (k+n) = m - (n::nat)"; | 
| 3339 | 471 | by (induct_tac "k" 1); | 
| 3234 | 472 | by (ALLGOALS Asm_simp_tac); | 
| 473 | qed "diff_cancel"; | |
| 474 | Addsimps [diff_cancel]; | |
| 475 | ||
| 5429 | 476 | Goal "(m+k) - (n+k) = m - (n::nat)"; | 
| 3234 | 477 | val add_commute_k = read_instantiate [("n","k")] add_commute;
 | 
| 5537 | 478 | by (asm_simp_tac (simpset() addsimps [add_commute_k]) 1); | 
| 3234 | 479 | qed "diff_cancel2"; | 
| 480 | Addsimps [diff_cancel2]; | |
| 481 | ||
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 482 | (*From Clemens Ballarin, proof by lcp*) | 
| 5429 | 483 | Goal "[| k<=n; n<=m |] ==> (m-k) - (n-k) = m-(n::nat)"; | 
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 484 | by (REPEAT (etac rev_mp 1)); | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 485 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 486 | by (ALLGOALS Asm_simp_tac); | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 487 | (*a confluence problem*) | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 488 | by (asm_simp_tac (simpset() addsimps [Suc_diff_le, le_Suc_eq]) 1); | 
| 3234 | 489 | qed "diff_right_cancel"; | 
| 490 | ||
| 5429 | 491 | Goal "n - (n+m) = 0"; | 
| 3339 | 492 | by (induct_tac "n" 1); | 
| 3234 | 493 | by (ALLGOALS Asm_simp_tac); | 
| 494 | qed "diff_add_0"; | |
| 495 | Addsimps [diff_add_0]; | |
| 496 | ||
| 5409 | 497 | |
| 3234 | 498 | (** Difference distributes over multiplication **) | 
| 499 | ||
| 5069 | 500 | Goal "!!m::nat. (m - n) * k = (m * k) - (n * k)"; | 
| 3234 | 501 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 502 | by (ALLGOALS Asm_simp_tac); | |
| 503 | qed "diff_mult_distrib" ; | |
| 504 | ||
| 5069 | 505 | Goal "!!m::nat. k * (m - n) = (k * m) - (k * n)"; | 
| 3234 | 506 | val mult_commute_k = read_instantiate [("m","k")] mult_commute;
 | 
| 4089 | 507 | by (simp_tac (simpset() addsimps [diff_mult_distrib, mult_commute_k]) 1); | 
| 3234 | 508 | qed "diff_mult_distrib2" ; | 
| 509 | (*NOT added as rewrites, since sometimes they are used from right-to-left*) | |
| 510 | ||
| 511 | ||
| 1713 | 512 | (*** Monotonicity of Multiplication ***) | 
| 513 | ||
| 5429 | 514 | Goal "i <= (j::nat) ==> i*k<=j*k"; | 
| 3339 | 515 | by (induct_tac "k" 1); | 
| 4089 | 516 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono]))); | 
| 1713 | 517 | qed "mult_le_mono1"; | 
| 518 | ||
| 519 | (*<=monotonicity, BOTH arguments*) | |
| 5429 | 520 | Goal "[| i <= (j::nat); k <= l |] ==> i*k <= j*l"; | 
| 2007 | 521 | by (etac (mult_le_mono1 RS le_trans) 1); | 
| 1713 | 522 | by (rtac le_trans 1); | 
| 2007 | 523 | by (stac mult_commute 2); | 
| 524 | by (etac mult_le_mono1 2); | |
| 4089 | 525 | by (simp_tac (simpset() addsimps [mult_commute]) 1); | 
| 1713 | 526 | qed "mult_le_mono"; | 
| 527 | ||
| 528 | (*strict, in 1st argument; proof is by induction on k>0*) | |
| 5429 | 529 | Goal "[| i<j; 0<k |] ==> k*i < k*j"; | 
| 5078 | 530 | by (eres_inst_tac [("m1","0")] (less_eq_Suc_add RS exE) 1);
 | 
| 1713 | 531 | by (Asm_simp_tac 1); | 
| 3339 | 532 | by (induct_tac "x" 1); | 
| 4089 | 533 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_less_mono]))); | 
| 1713 | 534 | qed "mult_less_mono2"; | 
| 535 | ||
| 5429 | 536 | Goal "[| i<j; 0<k |] ==> i*k < j*k"; | 
| 3457 | 537 | by (dtac mult_less_mono2 1); | 
| 4089 | 538 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [mult_commute]))); | 
| 3234 | 539 | qed "mult_less_mono1"; | 
| 540 | ||
| 5069 | 541 | Goal "(0 < m*n) = (0<m & 0<n)"; | 
| 3339 | 542 | by (induct_tac "m" 1); | 
| 543 | by (induct_tac "n" 2); | |
| 1713 | 544 | by (ALLGOALS Asm_simp_tac); | 
| 545 | qed "zero_less_mult_iff"; | |
| 4356 | 546 | Addsimps [zero_less_mult_iff]; | 
| 1713 | 547 | |
| 5069 | 548 | Goal "(m*n = 1) = (m=1 & n=1)"; | 
| 3339 | 549 | by (induct_tac "m" 1); | 
| 1795 | 550 | by (Simp_tac 1); | 
| 3339 | 551 | by (induct_tac "n" 1); | 
| 1795 | 552 | by (Simp_tac 1); | 
| 4089 | 553 | by (fast_tac (claset() addss simpset()) 1); | 
| 1795 | 554 | qed "mult_eq_1_iff"; | 
| 4356 | 555 | Addsimps [mult_eq_1_iff]; | 
| 1795 | 556 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 557 | Goal "0<k ==> (m*k < n*k) = (m<n)"; | 
| 4089 | 558 | by (safe_tac (claset() addSIs [mult_less_mono1])); | 
| 3234 | 559 | by (cut_facts_tac [less_linear] 1); | 
| 4389 | 560 | by (blast_tac (claset() addIs [mult_less_mono1] addEs [less_asym]) 1); | 
| 3234 | 561 | qed "mult_less_cancel2"; | 
| 562 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 563 | Goal "0<k ==> (k*m < k*n) = (m<n)"; | 
| 3457 | 564 | by (dtac mult_less_cancel2 1); | 
| 4089 | 565 | by (asm_full_simp_tac (simpset() addsimps [mult_commute]) 1); | 
| 3234 | 566 | qed "mult_less_cancel1"; | 
| 567 | Addsimps [mult_less_cancel1, mult_less_cancel2]; | |
| 568 | ||
| 5069 | 569 | Goal "(Suc k * m < Suc k * n) = (m < n)"; | 
| 4423 | 570 | by (rtac mult_less_cancel1 1); | 
| 4297 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 571 | by (Simp_tac 1); | 
| 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 572 | qed "Suc_mult_less_cancel1"; | 
| 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 573 | |
| 5069 | 574 | Goalw [le_def] "(Suc k * m <= Suc k * n) = (m <= n)"; | 
| 4297 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 575 | by (simp_tac (simpset_of HOL.thy) 1); | 
| 4423 | 576 | by (rtac Suc_mult_less_cancel1 1); | 
| 4297 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 577 | qed "Suc_mult_le_cancel1"; | 
| 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 578 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 579 | Goal "0<k ==> (m*k = n*k) = (m=n)"; | 
| 3234 | 580 | by (cut_facts_tac [less_linear] 1); | 
| 3724 | 581 | by Safe_tac; | 
| 3457 | 582 | by (assume_tac 2); | 
| 3234 | 583 | by (ALLGOALS (dtac mult_less_mono1 THEN' assume_tac)); | 
| 584 | by (ALLGOALS Asm_full_simp_tac); | |
| 585 | qed "mult_cancel2"; | |
| 586 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 587 | Goal "0<k ==> (k*m = k*n) = (m=n)"; | 
| 3457 | 588 | by (dtac mult_cancel2 1); | 
| 4089 | 589 | by (asm_full_simp_tac (simpset() addsimps [mult_commute]) 1); | 
| 3234 | 590 | qed "mult_cancel1"; | 
| 591 | Addsimps [mult_cancel1, mult_cancel2]; | |
| 592 | ||
| 5069 | 593 | Goal "(Suc k * m = Suc k * n) = (m = n)"; | 
| 4423 | 594 | by (rtac mult_cancel1 1); | 
| 4297 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 595 | by (Simp_tac 1); | 
| 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 596 | qed "Suc_mult_cancel1"; | 
| 
5defc2105cc8
added Suc_mult_less_cancel1, Suc_mult_le_cancel1, Suc_mult_cancel1;
 wenzelm parents: 
4158diff
changeset | 597 | |
| 3234 | 598 | |
| 1795 | 599 | (** Lemma for gcd **) | 
| 600 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 601 | Goal "m = m*n ==> n=1 | m=0"; | 
| 1795 | 602 | by (dtac sym 1); | 
| 603 | by (rtac disjCI 1); | |
| 604 | by (rtac nat_less_cases 1 THEN assume_tac 2); | |
| 4089 | 605 | by (fast_tac (claset() addSEs [less_SucE] addss simpset()) 1); | 
| 4356 | 606 | by (best_tac (claset() addDs [mult_less_mono2] addss simpset()) 1); | 
| 1795 | 607 | qed "mult_eq_self_implies_10"; | 
| 608 | ||
| 609 | ||
| 5983 | 610 | |
| 611 | ||
| 612 | (*---------------------------------------------------------------------------*) | |
| 613 | (* Various arithmetic proof procedures *) | |
| 614 | (*---------------------------------------------------------------------------*) | |
| 615 | ||
| 616 | (*---------------------------------------------------------------------------*) | |
| 617 | (* 1. Cancellation of common terms *) | |
| 618 | (*---------------------------------------------------------------------------*) | |
| 619 | ||
| 620 | (* Title: HOL/arith_data.ML | |
| 621 | ID: $Id$ | |
| 622 | Author: Markus Wenzel and Stefan Berghofer, TU Muenchen | |
| 623 | ||
| 624 | Setup various arithmetic proof procedures. | |
| 625 | *) | |
| 626 | ||
| 627 | signature ARITH_DATA = | |
| 628 | sig | |
| 6055 | 629 | val nat_cancel_sums_add: simproc list | 
| 5983 | 630 | val nat_cancel_sums: simproc list | 
| 631 | val nat_cancel_factor: simproc list | |
| 632 | val nat_cancel: simproc list | |
| 633 | end; | |
| 634 | ||
| 635 | structure ArithData: ARITH_DATA = | |
| 636 | struct | |
| 637 | ||
| 638 | ||
| 639 | (** abstract syntax of structure nat: 0, Suc, + **) | |
| 640 | ||
| 641 | (* mk_sum, mk_norm_sum *) | |
| 642 | ||
| 643 | val one = HOLogic.mk_nat 1; | |
| 644 | val mk_plus = HOLogic.mk_binop "op +"; | |
| 645 | ||
| 646 | fun mk_sum [] = HOLogic.zero | |
| 647 | | mk_sum [t] = t | |
| 648 | | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | |
| 649 | ||
| 650 | (*normal form of sums: Suc (... (Suc (a + (b + ...))))*) | |
| 651 | fun mk_norm_sum ts = | |
| 652 | let val (ones, sums) = partition (equal one) ts in | |
| 653 | funpow (length ones) HOLogic.mk_Suc (mk_sum sums) | |
| 654 | end; | |
| 655 | ||
| 656 | ||
| 657 | (* dest_sum *) | |
| 658 | ||
| 659 | val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT; | |
| 660 | ||
| 661 | fun dest_sum tm = | |
| 662 | if HOLogic.is_zero tm then [] | |
| 663 | else | |
| 664 | (case try HOLogic.dest_Suc tm of | |
| 665 | Some t => one :: dest_sum t | |
| 666 | | None => | |
| 667 | (case try dest_plus tm of | |
| 668 | Some (t, u) => dest_sum t @ dest_sum u | |
| 669 | | None => [tm])); | |
| 670 | ||
| 671 | ||
| 672 | (** generic proof tools **) | |
| 673 | ||
| 674 | (* prove conversions *) | |
| 675 | ||
| 676 | val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq; | |
| 677 | ||
| 678 | fun prove_conv expand_tac norm_tac sg (t, u) = | |
| 679 | mk_meta_eq (prove_goalw_cterm_nocheck [] (cterm_of sg (mk_eqv (t, u))) | |
| 680 | (K [expand_tac, norm_tac])) | |
| 681 |   handle ERROR => error ("The error(s) above occurred while trying to prove " ^
 | |
| 682 | (string_of_cterm (cterm_of sg (mk_eqv (t, u))))); | |
| 683 | ||
| 684 | val subst_equals = prove_goal HOL.thy "[| t = s; u = t |] ==> u = s" | |
| 685 | (fn prems => [cut_facts_tac prems 1, SIMPSET' asm_simp_tac 1]); | |
| 686 | ||
| 687 | ||
| 688 | (* rewriting *) | |
| 689 | ||
| 690 | fun simp_all rules = ALLGOALS (simp_tac (HOL_ss addsimps rules)); | |
| 691 | ||
| 692 | val add_rules = [add_Suc, add_Suc_right, add_0, add_0_right]; | |
| 693 | val mult_rules = [mult_Suc, mult_Suc_right, mult_0, mult_0_right]; | |
| 694 | ||
| 695 | ||
| 696 | ||
| 697 | (** cancel common summands **) | |
| 698 | ||
| 699 | structure Sum = | |
| 700 | struct | |
| 701 | val mk_sum = mk_norm_sum; | |
| 702 | val dest_sum = dest_sum; | |
| 703 | val prove_conv = prove_conv; | |
| 704 | val norm_tac = simp_all add_rules THEN simp_all add_ac; | |
| 705 | end; | |
| 706 | ||
| 707 | fun gen_uncancel_tac rule ct = | |
| 708 | rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1; | |
| 709 | ||
| 710 | ||
| 711 | (* nat eq *) | |
| 712 | ||
| 713 | structure EqCancelSums = CancelSumsFun | |
| 714 | (struct | |
| 715 | open Sum; | |
| 716 | val mk_bal = HOLogic.mk_eq; | |
| 717 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT; | |
| 718 | val uncancel_tac = gen_uncancel_tac add_left_cancel; | |
| 719 | end); | |
| 720 | ||
| 721 | ||
| 722 | (* nat less *) | |
| 723 | ||
| 724 | structure LessCancelSums = CancelSumsFun | |
| 725 | (struct | |
| 726 | open Sum; | |
| 727 | val mk_bal = HOLogic.mk_binrel "op <"; | |
| 728 | val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT; | |
| 729 | val uncancel_tac = gen_uncancel_tac add_left_cancel_less; | |
| 730 | end); | |
| 731 | ||
| 732 | ||
| 733 | (* nat le *) | |
| 734 | ||
| 735 | structure LeCancelSums = CancelSumsFun | |
| 736 | (struct | |
| 737 | open Sum; | |
| 738 | val mk_bal = HOLogic.mk_binrel "op <="; | |
| 739 | val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT; | |
| 740 | val uncancel_tac = gen_uncancel_tac add_left_cancel_le; | |
| 741 | end); | |
| 742 | ||
| 743 | ||
| 744 | (* nat diff *) | |
| 745 | ||
| 746 | structure DiffCancelSums = CancelSumsFun | |
| 747 | (struct | |
| 748 | open Sum; | |
| 749 | val mk_bal = HOLogic.mk_binop "op -"; | |
| 750 | val dest_bal = HOLogic.dest_bin "op -" HOLogic.natT; | |
| 751 | val uncancel_tac = gen_uncancel_tac diff_cancel; | |
| 752 | end); | |
| 753 | ||
| 754 | ||
| 755 | ||
| 756 | (** cancel common factor **) | |
| 757 | ||
| 758 | structure Factor = | |
| 759 | struct | |
| 760 | val mk_sum = mk_norm_sum; | |
| 761 | val dest_sum = dest_sum; | |
| 762 | val prove_conv = prove_conv; | |
| 763 | val norm_tac = simp_all (add_rules @ mult_rules) THEN simp_all add_ac; | |
| 764 | end; | |
| 765 | ||
| 766 | fun mk_cnat n = cterm_of (sign_of Nat.thy) (HOLogic.mk_nat n); | |
| 767 | ||
| 768 | fun gen_multiply_tac rule k = | |
| 769 | if k > 0 then | |
| 770 | rtac (instantiate' [] [None, Some (mk_cnat (k - 1))] (rule RS subst_equals)) 1 | |
| 771 | else no_tac; | |
| 772 | ||
| 773 | ||
| 774 | (* nat eq *) | |
| 775 | ||
| 776 | structure EqCancelFactor = CancelFactorFun | |
| 777 | (struct | |
| 778 | open Factor; | |
| 779 | val mk_bal = HOLogic.mk_eq; | |
| 780 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT; | |
| 781 | val multiply_tac = gen_multiply_tac Suc_mult_cancel1; | |
| 782 | end); | |
| 783 | ||
| 784 | ||
| 785 | (* nat less *) | |
| 786 | ||
| 787 | structure LessCancelFactor = CancelFactorFun | |
| 788 | (struct | |
| 789 | open Factor; | |
| 790 | val mk_bal = HOLogic.mk_binrel "op <"; | |
| 791 | val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT; | |
| 792 | val multiply_tac = gen_multiply_tac Suc_mult_less_cancel1; | |
| 793 | end); | |
| 794 | ||
| 795 | ||
| 796 | (* nat le *) | |
| 797 | ||
| 798 | structure LeCancelFactor = CancelFactorFun | |
| 799 | (struct | |
| 800 | open Factor; | |
| 801 | val mk_bal = HOLogic.mk_binrel "op <="; | |
| 802 | val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT; | |
| 803 | val multiply_tac = gen_multiply_tac Suc_mult_le_cancel1; | |
| 804 | end); | |
| 805 | ||
| 806 | ||
| 807 | ||
| 808 | (** prepare nat_cancel simprocs **) | |
| 809 | ||
| 810 | fun prep_pat s = Thm.read_cterm (sign_of Arith.thy) (s, HOLogic.termTVar); | |
| 811 | val prep_pats = map prep_pat; | |
| 812 | ||
| 813 | fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc; | |
| 814 | ||
| 815 | val eq_pats = prep_pats ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"]; | |
| 816 | val less_pats = prep_pats ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"]; | |
| 817 | val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"]; | |
| 818 | val diff_pats = prep_pats ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"]; | |
| 819 | ||
| 6055 | 820 | val nat_cancel_sums_add = map prep_simproc | 
| 5983 | 821 |   [("nateq_cancel_sums", eq_pats, EqCancelSums.proc),
 | 
| 822 |    ("natless_cancel_sums", less_pats, LessCancelSums.proc),
 | |
| 6055 | 823 |    ("natle_cancel_sums", le_pats, LeCancelSums.proc)];
 | 
| 824 | ||
| 825 | val nat_cancel_sums = nat_cancel_sums_add @ | |
| 826 |   [prep_simproc("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
 | |
| 5983 | 827 | |
| 828 | val nat_cancel_factor = map prep_simproc | |
| 829 |   [("nateq_cancel_factor", eq_pats, EqCancelFactor.proc),
 | |
| 830 |    ("natless_cancel_factor", less_pats, LessCancelFactor.proc),
 | |
| 831 |    ("natle_cancel_factor", le_pats, LeCancelFactor.proc)];
 | |
| 832 | ||
| 833 | val nat_cancel = nat_cancel_factor @ nat_cancel_sums; | |
| 834 | ||
| 835 | ||
| 836 | end; | |
| 837 | ||
| 838 | open ArithData; | |
| 839 | ||
| 840 | Addsimprocs nat_cancel; | |
| 841 | ||
| 842 | (*---------------------------------------------------------------------------*) | |
| 843 | (* 2. Linear arithmetic *) | |
| 844 | (*---------------------------------------------------------------------------*) | |
| 845 | ||
| 6101 | 846 | (* Parameters data for general linear arithmetic functor *) | 
| 847 | ||
| 848 | structure LA_Logic: LIN_ARITH_LOGIC = | |
| 5983 | 849 | struct | 
| 850 | val ccontr = ccontr; | |
| 851 | val conjI = conjI; | |
| 6101 | 852 | val neqE = linorder_neqE; | 
| 5983 | 853 | val notI = notI; | 
| 854 | val sym = sym; | |
| 6109 | 855 | val not_lessD = linorder_not_less RS iffD1; | 
| 6128 | 856 | val not_leD = linorder_not_le RS iffD1; | 
| 5983 | 857 | |
| 6128 | 858 | |
| 859 | fun mk_Eq thm = (thm RS Eq_FalseI) handle _ => (thm RS Eq_TrueI); | |
| 860 | ||
| 6073 | 861 | val mk_Trueprop = HOLogic.mk_Trueprop; | 
| 862 | ||
| 6079 | 863 | fun neg_prop(TP$(Const("Not",_)$t)) = TP$t
 | 
| 864 |   | neg_prop(TP$t) = TP $ (Const("Not",HOLogic.boolT-->HOLogic.boolT)$t);
 | |
| 6073 | 865 | |
| 6101 | 866 | fun is_False thm = | 
| 867 | let val _ $ t = #prop(rep_thm thm) | |
| 868 |   in t = Const("False",HOLogic.boolT) end;
 | |
| 869 | ||
| 6128 | 870 | fun is_nat(t) = fastype_of1 t = HOLogic.natT; | 
| 871 | ||
| 872 | fun mk_nat_thm sg t = | |
| 873 |   let val ct = cterm_of sg t  and cn = cterm_of sg (Var(("n",0),HOLogic.natT))
 | |
| 874 | in instantiate ([],[(cn,ct)]) le0 end; | |
| 875 | ||
| 6101 | 876 | end; | 
| 877 | ||
| 6128 | 878 | structure Nat_LA_Data (* : LIN_ARITH_DATA *) = | 
| 6101 | 879 | struct | 
| 880 | ||
| 6128 | 881 | val lessD = [Suc_leI]; | 
| 6101 | 882 | |
| 6151 | 883 | val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT; | 
| 884 | ||
| 5983 | 885 | (* Turn term into list of summand * multiplicity plus a constant *) | 
| 886 | fun poly(Const("Suc",_)$t, (p,i)) = poly(t, (p,i+1))
 | |
| 6059 
aa00e235ea27
In Main: moved Bin to the left to preserve the solver in its simpset.
 nipkow parents: 
6055diff
changeset | 887 |   | poly(Const("op +",_) $ s $ t, pi) = poly(s,poly(t,pi))
 | 
| 5983 | 888 | | poly(t,(p,i)) = | 
| 6059 
aa00e235ea27
In Main: moved Bin to the left to preserve the solver in its simpset.
 nipkow parents: 
6055diff
changeset | 889 |       if t = Const("0",HOLogic.natT) then (p,i)
 | 
| 5983 | 890 | else (case assoc(p,t) of None => ((t,1)::p,i) | 
| 891 | | Some m => (overwrite(p,(t,m+1)), i)) | |
| 6151 | 892 | fun poly(t, pi as (p,i)) = | 
| 893 | if HOLogic.is_zero t then pi else | |
| 894 | (case try HOLogic.dest_Suc t of | |
| 895 | Some u => poly(u, (p,i+1)) | |
| 896 | | None => (case try dest_plus t of | |
| 897 | Some(r,s) => poly(r,poly(s,pi)) | |
| 898 | | None => (case assoc(p,t) of None => ((t,1)::p,i) | |
| 899 | | Some m => (overwrite(p,(t,m+1)), i)))) | |
| 5983 | 900 | |
| 6059 
aa00e235ea27
In Main: moved Bin to the left to preserve the solver in its simpset.
 nipkow parents: 
6055diff
changeset | 901 | fun nnb T = T = ([HOLogic.natT,HOLogic.natT] ---> HOLogic.boolT); | 
| 
aa00e235ea27
In Main: moved Bin to the left to preserve the solver in its simpset.
 nipkow parents: 
6055diff
changeset | 902 | |
| 6128 | 903 | fun decomp2(rel,lhs,rhs) = | 
| 5983 | 904 | let val (p,i) = poly(lhs,([],0)) and (q,j) = poly(rhs,([],0)) | 
| 905 | in case rel of | |
| 906 | "op <" => Some(p,i,"<",q,j) | |
| 907 | | "op <=" => Some(p,i,"<=",q,j) | |
| 908 | | "op =" => Some(p,i,"=",q,j) | |
| 909 | | _ => None | |
| 910 | end; | |
| 911 | ||
| 912 | fun negate(Some(x,i,rel,y,j)) = Some(x,i,"~"^rel,y,j) | |
| 913 | | negate None = None; | |
| 914 | ||
| 6128 | 915 | fun decomp1(T,xxx) = if nnb T then decomp2 xxx else None; | 
| 916 | ||
| 917 | fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp1(T,(rel,lhs,rhs)) | |
| 5983 | 918 |   | decomp(_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
 | 
| 6128 | 919 | negate(decomp1(T,(rel,lhs,rhs))) | 
| 5983 | 920 | | decomp _ = None | 
| 6055 | 921 | |
| 5983 | 922 | (* reduce contradictory <= to False. | 
| 923 | Most of the work is done by the cancel tactics. | |
| 924 | *) | |
| 6151 | 925 | val add_rules = [add_0,add_0_right,Zero_not_Suc,Suc_not_Zero,le_0_eq]; | 
| 5983 | 926 | |
| 927 | val cancel_sums_ss = HOL_basic_ss addsimps add_rules | |
| 6055 | 928 | addsimprocs nat_cancel_sums_add; | 
| 5983 | 929 | |
| 930 | val simp = simplify cancel_sums_ss; | |
| 931 | ||
| 932 | val add_mono_thms = map (fn s => prove_goal Arith.thy s | |
| 933 | (fn prems => [cut_facts_tac prems 1, | |
| 934 | blast_tac (claset() addIs [add_le_mono]) 1])) | |
| 935 | ["(i <= j) & (k <= l) ==> i + k <= j + (l::nat)", | |
| 6055 | 936 | "(i = j) & (k <= l) ==> i + k <= j + (l::nat)", | 
| 937 | "(i <= j) & (k = l) ==> i + k <= j + (l::nat)", | |
| 938 | "(i = j) & (k = l) ==> i + k = j + (l::nat)" | |
| 5983 | 939 | ]; | 
| 940 | ||
| 6128 | 941 | end; | 
| 6055 | 942 | |
| 6128 | 943 | structure LA_Data_Ref = | 
| 944 | struct | |
| 945 | val add_mono_thms = ref Nat_LA_Data.add_mono_thms | |
| 946 | val lessD = ref Nat_LA_Data.lessD | |
| 947 | val decomp = ref Nat_LA_Data.decomp | |
| 948 | val simp = ref Nat_LA_Data.simp | |
| 5983 | 949 | end; | 
| 950 | ||
| 6128 | 951 | structure Fast_Arith = | 
| 952 | Fast_Lin_Arith(structure LA_Logic=LA_Logic and LA_Data=LA_Data_Ref); | |
| 5983 | 953 | |
| 6128 | 954 | val fast_arith_tac = Fast_Arith.lin_arith_tac; | 
| 6073 | 955 | |
| 6128 | 956 | val nat_arith_simproc_pats = | 
| 957 | map (fn s => Thm.read_cterm (sign_of Arith.thy) (s, HOLogic.boolT)) | |
| 958 | ["(m::nat) < n","(m::nat) <= n", "(m::nat) = n"]; | |
| 5983 | 959 | |
| 6128 | 960 | val fast_nat_arith_simproc = mk_simproc "fast_nat_arith" nat_arith_simproc_pats | 
| 961 | Fast_Arith.lin_arith_prover; | |
| 6073 | 962 | |
| 963 | Addsimprocs [fast_nat_arith_simproc]; | |
| 964 | ||
| 965 | (* Because of fast_nat_arith_simproc, the arithmetic solver is really only | |
| 966 | useful to detect inconsistencies among the premises for subgoals which are | |
| 967 | *not* themselves (in)equalities, because the latter activate | |
| 968 | fast_nat_arith_simproc anyway. However, it seems cheaper to activate the | |
| 969 | solver all the time rather than add the additional check. *) | |
| 970 | ||
| 6128 | 971 | simpset_ref () := (simpset() addSolver Fast_Arith.cut_lin_arith_tac); | 
| 6055 | 972 | |
| 973 | (* Elimination of `-' on nat due to John Harrison *) | |
| 974 | Goal "P(a - b::nat) = (!d. (b = a + d --> P 0) & (a = b + d --> P d))"; | |
| 975 | by(case_tac "a <= b" 1); | |
| 976 | by(Auto_tac); | |
| 977 | by(eres_inst_tac [("x","b-a")] allE 1);
 | |
| 978 | by(Auto_tac); | |
| 979 | qed "nat_diff_split"; | |
| 980 | ||
| 981 | (* FIXME: K true should be replaced by a sensible test to speed things up | |
| 6157 | 982 | in case there are lots of irrelevant terms involved; | 
| 983 | elimination of min/max can be optimized: | |
| 984 | (max m n + k <= r) = (m+k <= r & n+k <= r) | |
| 985 | (l <= min m n + k) = (l <= m+k & l <= n+k) | |
| 6055 | 986 | *) | 
| 6128 | 987 | val arith_tac = | 
| 6157 | 988 | refute_tac (K true) (REPEAT o split_tac[nat_diff_split,split_min,split_max]) | 
| 989 | ((REPEAT_DETERM o etac linorder_neqE) THEN' fast_arith_tac); | |
| 6055 | 990 | |
| 5983 | 991 | (*---------------------------------------------------------------------------*) | 
| 992 | (* End of proof procedures. Now go and USE them! *) | |
| 993 | (*---------------------------------------------------------------------------*) | |
| 994 | ||
| 4736 | 995 | (*** Subtraction laws -- mostly from Clemens Ballarin ***) | 
| 3234 | 996 | |
| 5429 | 997 | Goal "[| a < (b::nat); c <= a |] ==> a-c < b-c"; | 
| 6128 | 998 | by(arith_tac 1); | 
| 3234 | 999 | qed "diff_less_mono"; | 
| 1000 | ||
| 5429 | 1001 | Goal "a+b < (c::nat) ==> a < c-b"; | 
| 6128 | 1002 | by(arith_tac 1); | 
| 3234 | 1003 | qed "add_less_imp_less_diff"; | 
| 1004 | ||
| 5427 | 1005 | Goal "(i < j-k) = (i+k < (j::nat))"; | 
| 6128 | 1006 | by(arith_tac 1); | 
| 5427 | 1007 | qed "less_diff_conv"; | 
| 1008 | ||
| 5497 | 1009 | Goal "(j-k <= (i::nat)) = (j <= i+k)"; | 
| 6128 | 1010 | by(arith_tac 1); | 
| 5485 | 1011 | qed "le_diff_conv"; | 
| 1012 | ||
| 5497 | 1013 | Goal "k <= j ==> (i <= j-k) = (i+k <= (j::nat))"; | 
| 6128 | 1014 | by(arith_tac 1); | 
| 5497 | 1015 | qed "le_diff_conv2"; | 
| 1016 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1017 | Goal "Suc i <= n ==> Suc (n - Suc i) = n - i"; | 
| 6128 | 1018 | by(arith_tac 1); | 
| 3234 | 1019 | qed "Suc_diff_Suc"; | 
| 1020 | ||
| 5429 | 1021 | Goal "i <= (n::nat) ==> n - (n - i) = i"; | 
| 6128 | 1022 | by(arith_tac 1); | 
| 3234 | 1023 | qed "diff_diff_cancel"; | 
| 3381 
2bac33ec2b0d
New theorems le_add_diff_inverse, le_add_diff_inverse2
 paulson parents: 
3366diff
changeset | 1024 | Addsimps [diff_diff_cancel]; | 
| 3234 | 1025 | |
| 5429 | 1026 | Goal "k <= (n::nat) ==> m <= n + m - k"; | 
| 6128 | 1027 | by(arith_tac 1); | 
| 3234 | 1028 | qed "le_add_diff"; | 
| 1029 | ||
| 6055 | 1030 | Goal "[| 0<k; j<i |] ==> j+k-i < k"; | 
| 6128 | 1031 | by(arith_tac 1); | 
| 6055 | 1032 | qed "add_diff_less"; | 
| 3234 | 1033 | |
| 5356 | 1034 | Goal "m-1 < n ==> m <= n"; | 
| 6128 | 1035 | by(arith_tac 1); | 
| 5356 | 1036 | qed "pred_less_imp_le"; | 
| 1037 | ||
| 1038 | Goal "j<=i ==> i - j < Suc i - j"; | |
| 6128 | 1039 | by(arith_tac 1); | 
| 5356 | 1040 | qed "diff_less_Suc_diff"; | 
| 1041 | ||
| 1042 | Goal "i - j <= Suc i - j"; | |
| 6128 | 1043 | by(arith_tac 1); | 
| 5356 | 1044 | qed "diff_le_Suc_diff"; | 
| 1045 | AddIffs [diff_le_Suc_diff]; | |
| 1046 | ||
| 1047 | Goal "n - Suc i <= n - i"; | |
| 6128 | 1048 | by(arith_tac 1); | 
| 5356 | 1049 | qed "diff_Suc_le_diff"; | 
| 1050 | AddIffs [diff_Suc_le_diff]; | |
| 1051 | ||
| 5409 | 1052 | Goal "0 < n ==> (m <= n-1) = (m<n)"; | 
| 6128 | 1053 | by(arith_tac 1); | 
| 5409 | 1054 | qed "le_pred_eq"; | 
| 1055 | ||
| 1056 | Goal "0 < n ==> (m-1 < n) = (m<=n)"; | |
| 6128 | 1057 | by(arith_tac 1); | 
| 5409 | 1058 | qed "less_pred_eq"; | 
| 1059 | ||
| 5414 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 1060 | (*In ordinary notation: if 0<n and n<=m then m-n < m *) | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 1061 | Goal "[| 0<n; ~ m<n |] ==> m - n < m"; | 
| 6128 | 1062 | 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: 
5409diff
changeset | 1063 | qed "diff_less"; | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 1064 | |
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 1065 | Goal "[| 0<n; n<=m |] ==> m - n < m"; | 
| 6128 | 1066 | 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: 
5409diff
changeset | 1067 | qed "le_diff_less"; | 
| 
8a458866637c
changed Suc_diff_n to Suc_diff_le, with premise n <= m instead of n < Suc(m)
 paulson parents: 
5409diff
changeset | 1068 | |
| 5356 | 1069 | |
| 4732 | 1070 | |
| 3484 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 1071 | (** (Anti)Monotonicity of subtraction -- by Stefan Merz **) | 
| 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 1072 | |
| 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 1073 | (* Monotonicity of subtraction in first argument *) | 
| 6055 | 1074 | Goal "m <= (n::nat) ==> (m-l) <= (n-l)"; | 
| 6128 | 1075 | by(arith_tac 1); | 
| 6055 | 1076 | qed "diff_le_mono"; | 
| 3484 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 1077 | |
| 5429 | 1078 | Goal "m <= (n::nat) ==> (l-n) <= (l-m)"; | 
| 6128 | 1079 | by(arith_tac 1); | 
| 6055 | 1080 | qed "diff_le_mono2"; | 
| 5983 | 1081 | |
| 1082 | ||
| 1083 | (*This proof requires natdiff_cancel_sums*) | |
| 6055 | 1084 | Goal "[| m < (n::nat); m<l |] ==> (l-n) < (l-m)"; | 
| 6128 | 1085 | by(arith_tac 1); | 
| 6055 | 1086 | qed "diff_less_mono2"; | 
| 5983 | 1087 | |
| 6055 | 1088 | Goal "[| m-n = 0; n-m = 0 |] ==> m=n"; | 
| 6128 | 1089 | by(arith_tac 1); | 
| 6055 | 1090 | qed "diffs0_imp_equal"; |