| author | paulson | 
| Thu, 17 Jun 1999 10:35:01 +0200 | |
| changeset 6833 | 15d6c121d75f | 
| parent 6115 | c70bce7deb0f | 
| child 7030 | 53934985426a | 
| permissions | -rw-r--r-- | 
| 2608 | 1 | (* Title: HOL/NatDef.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | |
| 4 | Copyright 1991 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 5069 | 7 | Goal "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
 | 
| 2608 | 8 | by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1)); | 
| 9 | qed "Nat_fun_mono"; | |
| 10 | ||
| 11 | val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski); | |
| 12 | ||
| 13 | (* Zero is a natural number -- this also justifies the type definition*) | |
| 5069 | 14 | Goal "Zero_Rep: Nat"; | 
| 2608 | 15 | by (stac Nat_unfold 1); | 
| 16 | by (rtac (singletonI RS UnI1) 1); | |
| 17 | qed "Zero_RepI"; | |
| 18 | ||
| 5316 | 19 | Goal "i: Nat ==> Suc_Rep(i) : Nat"; | 
| 2608 | 20 | by (stac Nat_unfold 1); | 
| 21 | by (rtac (imageI RS UnI2) 1); | |
| 5316 | 22 | by (assume_tac 1); | 
| 2608 | 23 | qed "Suc_RepI"; | 
| 24 | ||
| 25 | (*** Induction ***) | |
| 26 | ||
| 5316 | 27 | val major::prems = Goal | 
| 2608 | 28 | "[| i: Nat; P(Zero_Rep); \ | 
| 29 | \ !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |] ==> P(i)"; | |
| 30 | by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1); | |
| 4089 | 31 | by (blast_tac (claset() addIs prems) 1); | 
| 2608 | 32 | qed "Nat_induct"; | 
| 33 | ||
| 5316 | 34 | val prems = Goalw [Zero_def,Suc_def] | 
| 2608 | 35 | "[| P(0); \ | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3023diff
changeset | 36 | \ !!n. P(n) ==> P(Suc(n)) |] ==> P(n)"; | 
| 2608 | 37 | by (rtac (Rep_Nat_inverse RS subst) 1); (*types force good instantiation*) | 
| 38 | by (rtac (Rep_Nat RS Nat_induct) 1); | |
| 39 | by (REPEAT (ares_tac prems 1 | |
| 40 | ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1)); | |
| 41 | qed "nat_induct"; | |
| 42 | ||
| 43 | (*Perform induction on n. *) | |
| 5187 
55f07169cf5f
Removed nat_case, nat_rec, and natE (now provided by datatype
 berghofe parents: 
5148diff
changeset | 44 | fun nat_ind_tac a i = | 
| 
55f07169cf5f
Removed nat_case, nat_rec, and natE (now provided by datatype
 berghofe parents: 
5148diff
changeset | 45 |   res_inst_tac [("n",a)] nat_induct i  THEN  rename_last_tac a [""] (i+1);
 | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3023diff
changeset | 46 | |
| 2608 | 47 | (*A special form of induction for reasoning about m<n and m-n*) | 
| 5316 | 48 | val prems = Goal | 
| 2608 | 49 | "[| !!x. P x 0; \ | 
| 50 | \ !!y. P 0 (Suc y); \ | |
| 51 | \ !!x y. [| P x y |] ==> P (Suc x) (Suc y) \ | |
| 52 | \ |] ==> P m n"; | |
| 53 | by (res_inst_tac [("x","m")] spec 1);
 | |
| 54 | by (nat_ind_tac "n" 1); | |
| 55 | by (rtac allI 2); | |
| 56 | by (nat_ind_tac "x" 2); | |
| 57 | by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1)); | |
| 58 | qed "diff_induct"; | |
| 59 | ||
| 60 | (*** Isomorphisms: Abs_Nat and Rep_Nat ***) | |
| 61 | ||
| 62 | (*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat), | |
| 63 | since we assume the isomorphism equations will one day be given by Isabelle*) | |
| 64 | ||
| 5069 | 65 | Goal "inj(Rep_Nat)"; | 
| 2608 | 66 | by (rtac inj_inverseI 1); | 
| 67 | by (rtac Rep_Nat_inverse 1); | |
| 68 | qed "inj_Rep_Nat"; | |
| 69 | ||
| 5069 | 70 | Goal "inj_on Abs_Nat Nat"; | 
| 4830 | 71 | by (rtac inj_on_inverseI 1); | 
| 2608 | 72 | by (etac Abs_Nat_inverse 1); | 
| 4830 | 73 | qed "inj_on_Abs_Nat"; | 
| 2608 | 74 | |
| 75 | (*** Distinctness of constructors ***) | |
| 76 | ||
| 5069 | 77 | Goalw [Zero_def,Suc_def] "Suc(m) ~= 0"; | 
| 4830 | 78 | by (rtac (inj_on_Abs_Nat RS inj_on_contraD) 1); | 
| 2608 | 79 | by (rtac Suc_Rep_not_Zero_Rep 1); | 
| 80 | by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1)); | |
| 81 | qed "Suc_not_Zero"; | |
| 82 | ||
| 83 | bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym);
 | |
| 84 | ||
| 85 | AddIffs [Suc_not_Zero,Zero_not_Suc]; | |
| 86 | ||
| 87 | bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
 | |
| 88 | val Zero_neq_Suc = sym RS Suc_neq_Zero; | |
| 89 | ||
| 90 | (** Injectiveness of Suc **) | |
| 91 | ||
| 5069 | 92 | Goalw [Suc_def] "inj(Suc)"; | 
| 2608 | 93 | by (rtac injI 1); | 
| 4830 | 94 | by (dtac (inj_on_Abs_Nat RS inj_onD) 1); | 
| 2608 | 95 | by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1)); | 
| 96 | by (dtac (inj_Suc_Rep RS injD) 1); | |
| 97 | by (etac (inj_Rep_Nat RS injD) 1); | |
| 98 | qed "inj_Suc"; | |
| 99 | ||
| 100 | val Suc_inject = inj_Suc RS injD; | |
| 101 | ||
| 5069 | 102 | Goal "(Suc(m)=Suc(n)) = (m=n)"; | 
| 2608 | 103 | by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); | 
| 104 | qed "Suc_Suc_eq"; | |
| 105 | ||
| 106 | AddIffs [Suc_Suc_eq]; | |
| 107 | ||
| 5069 | 108 | Goal "n ~= Suc(n)"; | 
| 2608 | 109 | by (nat_ind_tac "n" 1); | 
| 110 | by (ALLGOALS Asm_simp_tac); | |
| 111 | qed "n_not_Suc_n"; | |
| 112 | ||
| 113 | bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
 | |
| 114 | ||
| 5187 
55f07169cf5f
Removed nat_case, nat_rec, and natE (now provided by datatype
 berghofe parents: 
5148diff
changeset | 115 | (*** Basic properties of "less than" ***) | 
| 2608 | 116 | |
| 5069 | 117 | Goalw [wf_def, pred_nat_def] "wf(pred_nat)"; | 
| 3718 | 118 | by (Clarify_tac 1); | 
| 2608 | 119 | by (nat_ind_tac "x" 1); | 
| 3236 | 120 | by (ALLGOALS Blast_tac); | 
| 2608 | 121 | qed "wf_pred_nat"; | 
| 122 | ||
| 3378 | 123 | (*Used in TFL/post.sml*) | 
| 5069 | 124 | Goalw [less_def] "(m,n) : pred_nat^+ = (m<n)"; | 
| 3378 | 125 | by (rtac refl 1); | 
| 126 | qed "less_eq"; | |
| 127 | ||
| 2608 | 128 | (** Introduction properties **) | 
| 129 | ||
| 5316 | 130 | Goalw [less_def] "[| i<j; j<k |] ==> i<(k::nat)"; | 
| 2608 | 131 | by (rtac (trans_trancl RS transD) 1); | 
| 5316 | 132 | by (assume_tac 1); | 
| 133 | by (assume_tac 1); | |
| 2608 | 134 | qed "less_trans"; | 
| 135 | ||
| 5069 | 136 | Goalw [less_def, pred_nat_def] "n < Suc(n)"; | 
| 4089 | 137 | by (simp_tac (simpset() addsimps [r_into_trancl]) 1); | 
| 2608 | 138 | qed "lessI"; | 
| 139 | AddIffs [lessI]; | |
| 140 | ||
| 141 | (* i<j ==> i<Suc(j) *) | |
| 142 | bind_thm("less_SucI", lessI RSN (2, less_trans));
 | |
| 143 | Addsimps [less_SucI]; | |
| 144 | ||
| 5069 | 145 | Goal "0 < Suc(n)"; | 
| 2608 | 146 | by (nat_ind_tac "n" 1); | 
| 147 | by (rtac lessI 1); | |
| 148 | by (etac less_trans 1); | |
| 149 | by (rtac lessI 1); | |
| 150 | qed "zero_less_Suc"; | |
| 151 | AddIffs [zero_less_Suc]; | |
| 152 | ||
| 153 | (** Elimination properties **) | |
| 154 | ||
| 5316 | 155 | Goalw [less_def] "n<m ==> ~ m<(n::nat)"; | 
| 156 | by (blast_tac (claset() addIs [wf_pred_nat, wf_trancl RS wf_asym])1); | |
| 2608 | 157 | qed "less_not_sym"; | 
| 158 | ||
| 5474 
a2109bb8ce2b
well-formed asym rules; also adds less_irrefl, le_refl since order_refl
 paulson parents: 
5354diff
changeset | 159 | (* [| n<m; ~P ==> m<n |] ==> P *) | 
| 
a2109bb8ce2b
well-formed asym rules; also adds less_irrefl, le_refl since order_refl
 paulson parents: 
5354diff
changeset | 160 | bind_thm ("less_asym", less_not_sym RS swap);
 | 
| 2608 | 161 | |
| 5069 | 162 | Goalw [less_def] "~ n<(n::nat)"; | 
| 2608 | 163 | by (rtac notI 1); | 
| 164 | by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1); | |
| 165 | qed "less_not_refl"; | |
| 166 | ||
| 167 | (* n<n ==> R *) | |
| 168 | bind_thm ("less_irrefl", (less_not_refl RS notE));
 | |
| 5474 
a2109bb8ce2b
well-formed asym rules; also adds less_irrefl, le_refl since order_refl
 paulson parents: 
5354diff
changeset | 169 | AddSEs [less_irrefl]; | 
| 2608 | 170 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 171 | Goal "n<m ==> m ~= (n::nat)"; | 
| 5474 
a2109bb8ce2b
well-formed asym rules; also adds less_irrefl, le_refl since order_refl
 paulson parents: 
5354diff
changeset | 172 | by (Blast_tac 1); | 
| 2608 | 173 | qed "less_not_refl2"; | 
| 174 | ||
| 5354 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 175 | (* s < t ==> s ~= t *) | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 176 | bind_thm ("less_not_refl3", less_not_refl2 RS not_sym);
 | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 177 | |
| 2608 | 178 | |
| 5316 | 179 | val major::prems = Goalw [less_def, pred_nat_def] | 
| 2608 | 180 | "[| i<k; k=Suc(i) ==> P; !!j. [| i<j; k=Suc(j) |] ==> P \ | 
| 181 | \ |] ==> P"; | |
| 182 | by (rtac (major RS tranclE) 1); | |
| 3236 | 183 | by (ALLGOALS Full_simp_tac); | 
| 2608 | 184 | by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE' | 
| 3236 | 185 | eresolve_tac (prems@[asm_rl, Pair_inject]))); | 
| 2608 | 186 | qed "lessE"; | 
| 187 | ||
| 5069 | 188 | Goal "~ n<0"; | 
| 2608 | 189 | by (rtac notI 1); | 
| 190 | by (etac lessE 1); | |
| 191 | by (etac Zero_neq_Suc 1); | |
| 192 | by (etac Zero_neq_Suc 1); | |
| 193 | qed "not_less0"; | |
| 194 | ||
| 195 | AddIffs [not_less0]; | |
| 196 | ||
| 197 | (* n<0 ==> R *) | |
| 198 | bind_thm ("less_zeroE", not_less0 RS notE);
 | |
| 199 | ||
| 5316 | 200 | val [major,less,eq] = Goal | 
| 2608 | 201 | "[| m < Suc(n); m<n ==> P; m=n ==> P |] ==> P"; | 
| 202 | by (rtac (major RS lessE) 1); | |
| 203 | by (rtac eq 1); | |
| 2891 | 204 | by (Blast_tac 1); | 
| 2608 | 205 | by (rtac less 1); | 
| 2891 | 206 | by (Blast_tac 1); | 
| 2608 | 207 | qed "less_SucE"; | 
| 208 | ||
| 5069 | 209 | Goal "(m < Suc(n)) = (m < n | m = n)"; | 
| 4089 | 210 | by (blast_tac (claset() addSEs [less_SucE] addIs [less_trans]) 1); | 
| 2608 | 211 | qed "less_Suc_eq"; | 
| 212 | ||
| 5069 | 213 | Goal "(n<1) = (n=0)"; | 
| 4089 | 214 | by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); | 
| 3484 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 215 | qed "less_one"; | 
| 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 216 | AddIffs [less_one]; | 
| 
1e93eb09ebb9
Added the following lemmas tp Divides and a few others to Arith and NatDef:
 nipkow parents: 
3457diff
changeset | 217 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 218 | Goal "m<n ==> Suc(m) < Suc(n)"; | 
| 2608 | 219 | by (etac rev_mp 1); | 
| 220 | by (nat_ind_tac "n" 1); | |
| 5474 
a2109bb8ce2b
well-formed asym rules; also adds less_irrefl, le_refl since order_refl
 paulson parents: 
5354diff
changeset | 221 | by (ALLGOALS (fast_tac (claset() addEs [less_trans, lessE]))); | 
| 2608 | 222 | qed "Suc_mono"; | 
| 223 | ||
| 224 | (*"Less than" is a linear ordering*) | |
| 5069 | 225 | Goal "m<n | m=n | n<(m::nat)"; | 
| 2608 | 226 | by (nat_ind_tac "m" 1); | 
| 227 | by (nat_ind_tac "n" 1); | |
| 228 | by (rtac (refl RS disjI1 RS disjI2) 1); | |
| 229 | by (rtac (zero_less_Suc RS disjI1) 1); | |
| 4089 | 230 | by (blast_tac (claset() addIs [Suc_mono, less_SucI] addEs [lessE]) 1); | 
| 2608 | 231 | qed "less_linear"; | 
| 232 | ||
| 5069 | 233 | Goal "!!m::nat. (m ~= n) = (m<n | n<m)"; | 
| 4737 | 234 | by (cut_facts_tac [less_linear] 1); | 
| 5500 | 235 | by (Blast_tac 1); | 
| 4737 | 236 | qed "nat_neq_iff"; | 
| 237 | ||
| 2608 | 238 | qed_goal "nat_less_cases" thy | 
| 239 | "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m" | |
| 2935 | 240 | ( fn [major,eqCase,lessCase] => | 
| 2608 | 241 | [ | 
| 2935 | 242 | (rtac (less_linear RS disjE) 1), | 
| 2608 | 243 | (etac disjE 2), | 
| 2935 | 244 | (etac lessCase 1), | 
| 245 | (etac (sym RS eqCase) 1), | |
| 246 | (etac major 1) | |
| 2608 | 247 | ]); | 
| 248 | ||
| 4745 | 249 | |
| 250 | (** Inductive (?) properties **) | |
| 251 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 252 | Goal "[| m<n; Suc m ~= n |] ==> Suc(m) < n"; | 
| 4745 | 253 | by (full_simp_tac (simpset() addsimps [nat_neq_iff]) 1); | 
| 254 | by (blast_tac (claset() addSEs [less_irrefl, less_SucE] addEs [less_asym]) 1); | |
| 255 | qed "Suc_lessI"; | |
| 256 | ||
| 5316 | 257 | Goal "Suc(m) < n ==> m<n"; | 
| 258 | by (etac rev_mp 1); | |
| 4745 | 259 | by (nat_ind_tac "n" 1); | 
| 260 | by (ALLGOALS (fast_tac (claset() addSIs [lessI RS less_SucI] | |
| 261 | addEs [less_trans, lessE]))); | |
| 262 | qed "Suc_lessD"; | |
| 263 | ||
| 5316 | 264 | val [major,minor] = Goal | 
| 4745 | 265 | "[| Suc(i)<k; !!j. [| i<j; k=Suc(j) |] ==> P \ | 
| 266 | \ |] ==> P"; | |
| 267 | by (rtac (major RS lessE) 1); | |
| 268 | by (etac (lessI RS minor) 1); | |
| 269 | by (etac (Suc_lessD RS minor) 1); | |
| 270 | by (assume_tac 1); | |
| 271 | qed "Suc_lessE"; | |
| 272 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 273 | Goal "Suc(m) < Suc(n) ==> m<n"; | 
| 4745 | 274 | by (blast_tac (claset() addEs [lessE, make_elim Suc_lessD]) 1); | 
| 275 | qed "Suc_less_SucD"; | |
| 276 | ||
| 277 | ||
| 5069 | 278 | Goal "(Suc(m) < Suc(n)) = (m<n)"; | 
| 4745 | 279 | by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]); | 
| 280 | qed "Suc_less_eq"; | |
| 281 | Addsimps [Suc_less_eq]; | |
| 282 | ||
| 6109 | 283 | (*Goal "~(Suc(n) < n)"; | 
| 4745 | 284 | by (blast_tac (claset() addEs [Suc_lessD RS less_irrefl]) 1); | 
| 285 | qed "not_Suc_n_less_n"; | |
| 6109 | 286 | Addsimps [not_Suc_n_less_n];*) | 
| 4745 | 287 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 288 | Goal "i<j ==> j<k --> Suc i < k"; | 
| 4745 | 289 | by (nat_ind_tac "k" 1); | 
| 290 | by (ALLGOALS (asm_simp_tac (simpset()))); | |
| 291 | by (asm_simp_tac (simpset() addsimps [less_Suc_eq]) 1); | |
| 292 | by (blast_tac (claset() addDs [Suc_lessD]) 1); | |
| 293 | qed_spec_mp "less_trans_Suc"; | |
| 294 | ||
| 2608 | 295 | (*Can be used with less_Suc_eq to get n=m | n<m *) | 
| 5069 | 296 | Goal "(~ m < n) = (n < Suc(m))"; | 
| 2608 | 297 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 298 | by (ALLGOALS Asm_simp_tac); | |
| 299 | qed "not_less_eq"; | |
| 300 | ||
| 301 | (*Complete induction, aka course-of-values induction*) | |
| 5316 | 302 | val prems = Goalw [less_def] | 
| 2608 | 303 | "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |] ==> P(n)"; | 
| 304 | by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); | |
| 305 | by (eresolve_tac prems 1); | |
| 306 | qed "less_induct"; | |
| 307 | ||
| 308 | (*** Properties of <= ***) | |
| 309 | ||
| 5500 | 310 | (*Was le_eq_less_Suc, but this orientation is more useful*) | 
| 311 | Goalw [le_def] "(m < Suc n) = (m <= n)"; | |
| 312 | by (rtac (not_less_eq RS sym) 1); | |
| 313 | qed "less_Suc_eq_le"; | |
| 2608 | 314 | |
| 3343 | 315 | (* m<=n ==> m < Suc n *) | 
| 5500 | 316 | bind_thm ("le_imp_less_Suc", less_Suc_eq_le RS iffD2);
 | 
| 3343 | 317 | |
| 5069 | 318 | Goalw [le_def] "0 <= n"; | 
| 2608 | 319 | by (rtac not_less0 1); | 
| 320 | qed "le0"; | |
| 6075 | 321 | AddIffs [le0]; | 
| 2608 | 322 | |
| 5069 | 323 | Goalw [le_def] "~ Suc n <= n"; | 
| 2608 | 324 | by (Simp_tac 1); | 
| 325 | qed "Suc_n_not_le_n"; | |
| 326 | ||
| 5069 | 327 | Goalw [le_def] "(i <= 0) = (i = 0)"; | 
| 2608 | 328 | by (nat_ind_tac "i" 1); | 
| 329 | by (ALLGOALS Asm_simp_tac); | |
| 330 | qed "le_0_eq"; | |
| 4614 | 331 | AddIffs [le_0_eq]; | 
| 2608 | 332 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 333 | Goal "(m <= Suc(n)) = (m<=n | m = Suc n)"; | 
| 5500 | 334 | by (simp_tac (simpset() delsimps [less_Suc_eq_le] | 
| 335 | addsimps [less_Suc_eq_le RS sym, less_Suc_eq]) 1); | |
| 3355 | 336 | qed "le_Suc_eq"; | 
| 337 | ||
| 4614 | 338 | (* [| m <= Suc n; m <= n ==> R; m = Suc n ==> R |] ==> R *) | 
| 339 | bind_thm ("le_SucE", le_Suc_eq RS iffD1 RS disjE);
 | |
| 340 | ||
| 5316 | 341 | Goalw [le_def] "~n<m ==> m<=(n::nat)"; | 
| 342 | by (assume_tac 1); | |
| 2608 | 343 | qed "leI"; | 
| 344 | ||
| 5316 | 345 | Goalw [le_def] "m<=n ==> ~ n < (m::nat)"; | 
| 346 | by (assume_tac 1); | |
| 2608 | 347 | qed "leD"; | 
| 348 | ||
| 349 | val leE = make_elim leD; | |
| 350 | ||
| 5069 | 351 | Goal "(~n<m) = (m<=(n::nat))"; | 
| 4089 | 352 | by (blast_tac (claset() addIs [leI] addEs [leE]) 1); | 
| 2608 | 353 | qed "not_less_iff_le"; | 
| 354 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 355 | Goalw [le_def] "~ m <= n ==> n<(m::nat)"; | 
| 2891 | 356 | by (Blast_tac 1); | 
| 2608 | 357 | qed "not_leE"; | 
| 358 | ||
| 5069 | 359 | Goalw [le_def] "(~n<=m) = (m<(n::nat))"; | 
| 4599 | 360 | by (Simp_tac 1); | 
| 361 | qed "not_le_iff_less"; | |
| 362 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 363 | Goalw [le_def] "m < n ==> Suc(m) <= n"; | 
| 4089 | 364 | by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); | 
| 365 | by (blast_tac (claset() addSEs [less_irrefl,less_asym]) 1); | |
| 3343 | 366 | qed "Suc_leI"; (*formerly called lessD*) | 
| 2608 | 367 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 368 | Goalw [le_def] "Suc(m) <= n ==> m <= n"; | 
| 4089 | 369 | by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1); | 
| 2608 | 370 | qed "Suc_leD"; | 
| 371 | ||
| 372 | (* stronger version of Suc_leD *) | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 373 | Goalw [le_def] "Suc m <= n ==> m < n"; | 
| 4089 | 374 | by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1); | 
| 2608 | 375 | by (cut_facts_tac [less_linear] 1); | 
| 2891 | 376 | by (Blast_tac 1); | 
| 2608 | 377 | qed "Suc_le_lessD"; | 
| 378 | ||
| 5069 | 379 | Goal "(Suc m <= n) = (m < n)"; | 
| 4089 | 380 | by (blast_tac (claset() addIs [Suc_leI, Suc_le_lessD]) 1); | 
| 2608 | 381 | qed "Suc_le_eq"; | 
| 382 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 383 | Goalw [le_def] "m <= n ==> m <= Suc n"; | 
| 4089 | 384 | by (blast_tac (claset() addDs [Suc_lessD]) 1); | 
| 2608 | 385 | qed "le_SucI"; | 
| 386 | Addsimps[le_SucI]; | |
| 387 | ||
| 6109 | 388 | (*bind_thm ("le_Suc", not_Suc_n_less_n RS leI);*)
 | 
| 2608 | 389 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 390 | Goalw [le_def] "m < n ==> m <= (n::nat)"; | 
| 4089 | 391 | by (blast_tac (claset() addEs [less_asym]) 1); | 
| 2608 | 392 | qed "less_imp_le"; | 
| 393 | ||
| 5591 | 394 | (*For instance, (Suc m < Suc n) = (Suc m <= n) = (m<n) *) | 
| 395 | val le_simps = [less_imp_le, less_Suc_eq_le, Suc_le_eq]; | |
| 396 | ||
| 5354 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 397 | |
| 3343 | 398 | (** Equivalence of m<=n and m<n | m=n **) | 
| 399 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 400 | Goalw [le_def] "m <= n ==> m < n | m=(n::nat)"; | 
| 2608 | 401 | by (cut_facts_tac [less_linear] 1); | 
| 4089 | 402 | by (blast_tac (claset() addEs [less_irrefl,less_asym]) 1); | 
| 2608 | 403 | qed "le_imp_less_or_eq"; | 
| 404 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 405 | Goalw [le_def] "m<n | m=n ==> m <=(n::nat)"; | 
| 2608 | 406 | by (cut_facts_tac [less_linear] 1); | 
| 4089 | 407 | by (blast_tac (claset() addSEs [less_irrefl] addEs [less_asym]) 1); | 
| 2608 | 408 | qed "less_or_eq_imp_le"; | 
| 409 | ||
| 5069 | 410 | Goal "(m <= (n::nat)) = (m < n | m=n)"; | 
| 2608 | 411 | by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1)); | 
| 412 | qed "le_eq_less_or_eq"; | |
| 413 | ||
| 4635 | 414 | (*Useful with Blast_tac. m=n ==> m<=n *) | 
| 415 | bind_thm ("eq_imp_le", disjI2 RS less_or_eq_imp_le);
 | |
| 416 | ||
| 5069 | 417 | Goal "n <= (n::nat)"; | 
| 4089 | 418 | by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); | 
| 2608 | 419 | qed "le_refl"; | 
| 420 | ||
| 5354 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 421 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 422 | Goal "[| i <= j; j < k |] ==> i < (k::nat)"; | 
| 4468 | 423 | by (blast_tac (claset() addSDs [le_imp_less_or_eq] | 
| 424 | addIs [less_trans]) 1); | |
| 2608 | 425 | qed "le_less_trans"; | 
| 426 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 427 | Goal "[| i < j; j <= k |] ==> i < (k::nat)"; | 
| 4468 | 428 | by (blast_tac (claset() addSDs [le_imp_less_or_eq] | 
| 429 | addIs [less_trans]) 1); | |
| 2608 | 430 | qed "less_le_trans"; | 
| 431 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 432 | Goal "[| i <= j; j <= k |] ==> i <= (k::nat)"; | 
| 4468 | 433 | by (blast_tac (claset() addSDs [le_imp_less_or_eq] | 
| 434 | addIs [less_or_eq_imp_le, less_trans]) 1); | |
| 2608 | 435 | qed "le_trans"; | 
| 436 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 437 | Goal "[| m <= n; n <= m |] ==> m = (n::nat)"; | 
| 4468 | 438 | (*order_less_irrefl could make this proof fail*) | 
| 439 | by (blast_tac (claset() addSDs [le_imp_less_or_eq] | |
| 440 | addSEs [less_irrefl] addEs [less_asym]) 1); | |
| 2608 | 441 | qed "le_anti_sym"; | 
| 442 | ||
| 5069 | 443 | Goal "(Suc(n) <= Suc(m)) = (n <= m)"; | 
| 5500 | 444 | by (simp_tac (simpset() addsimps le_simps) 1); | 
| 2608 | 445 | qed "Suc_le_mono"; | 
| 446 | ||
| 447 | AddIffs [Suc_le_mono]; | |
| 448 | ||
| 5500 | 449 | (* Axiom 'order_less_le' of class 'order': *) | 
| 5069 | 450 | Goal "(m::nat) < n = (m <= n & m ~= n)"; | 
| 4737 | 451 | by (simp_tac (simpset() addsimps [le_def, nat_neq_iff]) 1); | 
| 452 | by (blast_tac (claset() addSEs [less_asym]) 1); | |
| 2608 | 453 | qed "nat_less_le"; | 
| 454 | ||
| 5354 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 455 | (* [| m <= n; m ~= n |] ==> m < n *) | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 456 | bind_thm ("le_neq_implies_less", [nat_less_le, conjI] MRS iffD2);
 | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 457 | |
| 4640 | 458 | (* Axiom 'linorder_linear' of class 'linorder': *) | 
| 5069 | 459 | Goal "(m::nat) <= n | n <= m"; | 
| 4640 | 460 | by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); | 
| 461 | by (cut_facts_tac [less_linear] 1); | |
| 5132 | 462 | by (Blast_tac 1); | 
| 4640 | 463 | qed "nat_le_linear"; | 
| 464 | ||
| 5354 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 465 | Goal "~ n < m ==> (n < Suc m) = (n = m)"; | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 466 | by (blast_tac (claset() addSEs [less_SucE]) 1); | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 467 | qed "not_less_less_Suc_eq"; | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 468 | |
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 469 | |
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 470 | (*Rewrite (n < Suc m) to (n=m) if ~ n<m or m<=n hold. | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 471 | Not suitable as default simprules because they often lead to looping*) | 
| 
da63d9b35caf
new theorems; adds [le_refl, less_imp_le] as simprules
 paulson parents: 
5316diff
changeset | 472 | val not_less_simps = [not_less_less_Suc_eq, leD RS not_less_less_Suc_eq]; | 
| 4640 | 473 | |
| 2608 | 474 | (** LEAST -- the least number operator **) | 
| 475 | ||
| 5069 | 476 | Goal "(! m::nat. P m --> n <= m) = (! m. m < n --> ~ P m)"; | 
| 4089 | 477 | by (blast_tac (claset() addIs [leI] addEs [leE]) 1); | 
| 3143 
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
 nipkow parents: 
3085diff
changeset | 478 | val lemma = result(); | 
| 
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
 nipkow parents: 
3085diff
changeset | 479 | |
| 
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
 nipkow parents: 
3085diff
changeset | 480 | (* This is an old def of Least for nat, which is derived for compatibility *) | 
| 5069 | 481 | Goalw [Least_def] | 
| 3143 
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
 nipkow parents: 
3085diff
changeset | 482 | "(LEAST n::nat. P n) == (@n. P(n) & (ALL m. m < n --> ~P(m)))"; | 
| 4089 | 483 | by (simp_tac (simpset() addsimps [lemma]) 1); | 
| 3143 
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
 nipkow parents: 
3085diff
changeset | 484 | qed "Least_nat_def"; | 
| 
d60e49b86c6a
Modified def of Least, which, as Markus correctly complained, looked like
 nipkow parents: 
3085diff
changeset | 485 | |
| 5316 | 486 | val [prem1,prem2] = Goalw [Least_nat_def] | 
| 3842 | 487 | "[| P(k::nat); !!x. x<k ==> ~P(x) |] ==> (LEAST x. P(x)) = k"; | 
| 2608 | 488 | by (rtac select_equality 1); | 
| 4089 | 489 | by (blast_tac (claset() addSIs [prem1,prem2]) 1); | 
| 2608 | 490 | by (cut_facts_tac [less_linear] 1); | 
| 4089 | 491 | by (blast_tac (claset() addSIs [prem1] addSDs [prem2]) 1); | 
| 2608 | 492 | qed "Least_equality"; | 
| 493 | ||
| 5316 | 494 | Goal "P(k::nat) ==> P(LEAST x. P(x))"; | 
| 495 | by (etac rev_mp 1); | |
| 2608 | 496 | by (res_inst_tac [("n","k")] less_induct 1);
 | 
| 497 | by (rtac impI 1); | |
| 498 | by (rtac classical 1); | |
| 499 | by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
 | |
| 500 | by (assume_tac 1); | |
| 501 | by (assume_tac 2); | |
| 2891 | 502 | by (Blast_tac 1); | 
| 2608 | 503 | qed "LeastI"; | 
| 504 | ||
| 505 | (*Proof is almost identical to the one above!*) | |
| 5316 | 506 | Goal "P(k::nat) ==> (LEAST x. P(x)) <= k"; | 
| 507 | by (etac rev_mp 1); | |
| 2608 | 508 | by (res_inst_tac [("n","k")] less_induct 1);
 | 
| 509 | by (rtac impI 1); | |
| 510 | by (rtac classical 1); | |
| 511 | by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
 | |
| 512 | by (assume_tac 1); | |
| 513 | by (rtac le_refl 2); | |
| 4089 | 514 | by (blast_tac (claset() addIs [less_imp_le,le_trans]) 1); | 
| 2608 | 515 | qed "Least_le"; | 
| 516 | ||
| 5316 | 517 | Goal "k < (LEAST x. P(x)) ==> ~P(k::nat)"; | 
| 2608 | 518 | by (rtac notI 1); | 
| 5316 | 519 | by (etac (rewrite_rule [le_def] Least_le RS notE) 1 THEN assume_tac 1); | 
| 2608 | 520 | qed "not_less_Least"; | 
| 521 | ||
| 5983 | 522 | (* [| m ~= n; m < n ==> P; n < m ==> P |] ==> P *) | 
| 4737 | 523 | bind_thm("nat_neqE", nat_neq_iff RS iffD1 RS disjE);
 |