| author | clasohm | 
| Thu, 14 Mar 1996 12:21:07 +0100 | |
| changeset 1578 | b58a6182e184 | 
| parent 1552 | 6f71b5d46700 | 
| child 1618 | 372880456b5b | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/nat | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | 
| 923 | 4 | Copyright 1991 University of Cambridge | 
| 5 | ||
| 6 | For nat.thy. Type nat is defined as a set (Nat) over the type ind. | |
| 7 | *) | |
| 8 | ||
| 9 | open Nat; | |
| 10 | ||
| 11 | goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
 | |
| 12 | by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1)); | |
| 13 | qed "Nat_fun_mono"; | |
| 14 | ||
| 15 | val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski); | |
| 16 | ||
| 17 | (* Zero is a natural number -- this also justifies the type definition*) | |
| 18 | goal Nat.thy "Zero_Rep: Nat"; | |
| 19 | by (rtac (Nat_unfold RS ssubst) 1); | |
| 20 | by (rtac (singletonI RS UnI1) 1); | |
| 21 | qed "Zero_RepI"; | |
| 22 | ||
| 23 | val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat"; | |
| 24 | by (rtac (Nat_unfold RS ssubst) 1); | |
| 25 | by (rtac (imageI RS UnI2) 1); | |
| 26 | by (resolve_tac prems 1); | |
| 27 | qed "Suc_RepI"; | |
| 28 | ||
| 29 | (*** Induction ***) | |
| 30 | ||
| 31 | val major::prems = goal Nat.thy | |
| 32 | "[| i: Nat; P(Zero_Rep); \ | |
| 33 | \ !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |] ==> P(i)"; | |
| 34 | by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1); | |
| 35 | by (fast_tac (set_cs addIs prems) 1); | |
| 36 | qed "Nat_induct"; | |
| 37 | ||
| 38 | val prems = goalw Nat.thy [Zero_def,Suc_def] | |
| 39 | "[| P(0); \ | |
| 40 | \ !!k. P(k) ==> P(Suc(k)) |] ==> P(n)"; | |
| 41 | by (rtac (Rep_Nat_inverse RS subst) 1); (*types force good instantiation*) | |
| 42 | by (rtac (Rep_Nat RS Nat_induct) 1); | |
| 43 | by (REPEAT (ares_tac prems 1 | |
| 44 | ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1)); | |
| 45 | qed "nat_induct"; | |
| 46 | ||
| 47 | (*Perform induction on n. *) | |
| 48 | fun nat_ind_tac a i = | |
| 49 |     EVERY [res_inst_tac [("n",a)] nat_induct i,
 | |
| 1465 | 50 | rename_last_tac a ["1"] (i+1)]; | 
| 923 | 51 | |
| 52 | (*A special form of induction for reasoning about m<n and m-n*) | |
| 53 | val prems = goal Nat.thy | |
| 54 | "[| !!x. P x 0; \ | |
| 55 | \ !!y. P 0 (Suc y); \ | |
| 56 | \ !!x y. [| P x y |] ==> P (Suc x) (Suc y) \ | |
| 57 | \ |] ==> P m n"; | |
| 58 | by (res_inst_tac [("x","m")] spec 1);
 | |
| 59 | by (nat_ind_tac "n" 1); | |
| 60 | by (rtac allI 2); | |
| 61 | by (nat_ind_tac "x" 2); | |
| 62 | by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1)); | |
| 63 | qed "diff_induct"; | |
| 64 | ||
| 65 | (*Case analysis on the natural numbers*) | |
| 66 | val prems = goal Nat.thy | |
| 67 | "[| n=0 ==> P; !!x. n = Suc(x) ==> P |] ==> P"; | |
| 68 | by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1); | |
| 69 | by (fast_tac (HOL_cs addSEs prems) 1); | |
| 70 | by (nat_ind_tac "n" 1); | |
| 71 | by (rtac (refl RS disjI1) 1); | |
| 72 | by (fast_tac HOL_cs 1); | |
| 73 | qed "natE"; | |
| 74 | ||
| 75 | (*** Isomorphisms: Abs_Nat and Rep_Nat ***) | |
| 76 | ||
| 77 | (*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat), | |
| 78 | since we assume the isomorphism equations will one day be given by Isabelle*) | |
| 79 | ||
| 80 | goal Nat.thy "inj(Rep_Nat)"; | |
| 81 | by (rtac inj_inverseI 1); | |
| 82 | by (rtac Rep_Nat_inverse 1); | |
| 83 | qed "inj_Rep_Nat"; | |
| 84 | ||
| 85 | goal Nat.thy "inj_onto Abs_Nat Nat"; | |
| 86 | by (rtac inj_onto_inverseI 1); | |
| 87 | by (etac Abs_Nat_inverse 1); | |
| 88 | qed "inj_onto_Abs_Nat"; | |
| 89 | ||
| 90 | (*** Distinctness of constructors ***) | |
| 91 | ||
| 92 | goalw Nat.thy [Zero_def,Suc_def] "Suc(m) ~= 0"; | |
| 93 | by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1); | |
| 94 | by (rtac Suc_Rep_not_Zero_Rep 1); | |
| 95 | by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1)); | |
| 96 | qed "Suc_not_Zero"; | |
| 97 | ||
| 98 | bind_thm ("Zero_not_Suc", (Suc_not_Zero RS not_sym));
 | |
| 99 | ||
| 1301 | 100 | Addsimps [Suc_not_Zero,Zero_not_Suc]; | 
| 101 | ||
| 923 | 102 | bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
 | 
| 103 | val Zero_neq_Suc = sym RS Suc_neq_Zero; | |
| 104 | ||
| 105 | (** Injectiveness of Suc **) | |
| 106 | ||
| 107 | goalw Nat.thy [Suc_def] "inj(Suc)"; | |
| 108 | by (rtac injI 1); | |
| 109 | by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1); | |
| 110 | by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1)); | |
| 111 | by (dtac (inj_Suc_Rep RS injD) 1); | |
| 112 | by (etac (inj_Rep_Nat RS injD) 1); | |
| 113 | qed "inj_Suc"; | |
| 114 | ||
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 115 | val Suc_inject = inj_Suc RS injD; | 
| 923 | 116 | |
| 117 | goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)"; | |
| 118 | by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); | |
| 119 | qed "Suc_Suc_eq"; | |
| 120 | ||
| 121 | goal Nat.thy "n ~= Suc(n)"; | |
| 122 | by (nat_ind_tac "n" 1); | |
| 1301 | 123 | by (ALLGOALS(asm_simp_tac (!simpset addsimps [Suc_Suc_eq]))); | 
| 923 | 124 | qed "n_not_Suc_n"; | 
| 125 | ||
| 126 | val Suc_n_not_n = n_not_Suc_n RS not_sym; | |
| 127 | ||
| 128 | (*** nat_case -- the selection operator for nat ***) | |
| 129 | ||
| 130 | goalw Nat.thy [nat_case_def] "nat_case a f 0 = a"; | |
| 131 | by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1); | |
| 132 | qed "nat_case_0"; | |
| 133 | ||
| 134 | goalw Nat.thy [nat_case_def] "nat_case a f (Suc k) = f(k)"; | |
| 135 | by (fast_tac (set_cs addIs [select_equality] | |
| 1465 | 136 | addEs [make_elim Suc_inject, Suc_neq_Zero]) 1); | 
| 923 | 137 | qed "nat_case_Suc"; | 
| 138 | ||
| 139 | (** Introduction rules for 'pred_nat' **) | |
| 140 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
962diff
changeset | 141 | goalw Nat.thy [pred_nat_def] "(n, Suc(n)) : pred_nat"; | 
| 923 | 142 | by (fast_tac set_cs 1); | 
| 143 | qed "pred_natI"; | |
| 144 | ||
| 145 | val major::prems = goalw Nat.thy [pred_nat_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
962diff
changeset | 146 | "[| p : pred_nat; !!x n. [| p = (n, Suc(n)) |] ==> R \ | 
| 923 | 147 | \ |] ==> R"; | 
| 148 | by (rtac (major RS CollectE) 1); | |
| 149 | by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1)); | |
| 150 | qed "pred_natE"; | |
| 151 | ||
| 152 | goalw Nat.thy [wf_def] "wf(pred_nat)"; | |
| 153 | by (strip_tac 1); | |
| 154 | by (nat_ind_tac "x" 1); | |
| 155 | by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, | |
| 1465 | 156 | make_elim Suc_inject]) 2); | 
| 923 | 157 | by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, Zero_neq_Suc]) 1); | 
| 158 | qed "wf_pred_nat"; | |
| 159 | ||
| 160 | ||
| 161 | (*** nat_rec -- by wf recursion on pred_nat ***) | |
| 162 | ||
| 1475 | 163 | (* The unrolling rule for nat_rec *) | 
| 164 | goal Nat.thy | |
| 165 | "(%n. nat_rec n c d) = wfrec pred_nat (%f. nat_case ?c (%m. ?d m (f m)))"; | |
| 166 | by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1); | |
| 167 | bind_thm("nat_rec_unfold", wf_pred_nat RS 
 | |
| 168 | ((result() RS eq_reflection) RS def_wfrec)); | |
| 169 | ||
| 170 | (*--------------------------------------------------------------------------- | |
| 171 | * Old: | |
| 172 |  * bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); 
 | |
| 173 | *---------------------------------------------------------------------------*) | |
| 923 | 174 | |
| 175 | (** conversion rules **) | |
| 176 | ||
| 177 | goal Nat.thy "nat_rec 0 c h = c"; | |
| 178 | by (rtac (nat_rec_unfold RS trans) 1); | |
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 179 | by (simp_tac (!simpset addsimps [nat_case_0]) 1); | 
| 923 | 180 | qed "nat_rec_0"; | 
| 181 | ||
| 182 | goal Nat.thy "nat_rec (Suc n) c h = h n (nat_rec n c h)"; | |
| 183 | by (rtac (nat_rec_unfold RS trans) 1); | |
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 184 | by (simp_tac (!simpset addsimps [nat_case_Suc, pred_natI, cut_apply]) 1); | 
| 923 | 185 | qed "nat_rec_Suc"; | 
| 186 | ||
| 187 | (*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) | |
| 188 | val [rew] = goal Nat.thy | |
| 189 | "[| !!n. f(n) == nat_rec n c h |] ==> f(0) = c"; | |
| 190 | by (rewtac rew); | |
| 191 | by (rtac nat_rec_0 1); | |
| 192 | qed "def_nat_rec_0"; | |
| 193 | ||
| 194 | val [rew] = goal Nat.thy | |
| 195 | "[| !!n. f(n) == nat_rec n c h |] ==> f(Suc(n)) = h n (f n)"; | |
| 196 | by (rewtac rew); | |
| 197 | by (rtac nat_rec_Suc 1); | |
| 198 | qed "def_nat_rec_Suc"; | |
| 199 | ||
| 200 | fun nat_recs def = | |
| 201 | [standard (def RS def_nat_rec_0), | |
| 202 | standard (def RS def_nat_rec_Suc)]; | |
| 203 | ||
| 204 | ||
| 205 | (*** Basic properties of "less than" ***) | |
| 206 | ||
| 207 | (** Introduction properties **) | |
| 208 | ||
| 209 | val prems = goalw Nat.thy [less_def] "[| i<j; j<k |] ==> i<(k::nat)"; | |
| 210 | by (rtac (trans_trancl RS transD) 1); | |
| 211 | by (resolve_tac prems 1); | |
| 212 | by (resolve_tac prems 1); | |
| 213 | qed "less_trans"; | |
| 214 | ||
| 215 | goalw Nat.thy [less_def] "n < Suc(n)"; | |
| 216 | by (rtac (pred_natI RS r_into_trancl) 1); | |
| 217 | qed "lessI"; | |
| 1301 | 218 | Addsimps [lessI]; | 
| 923 | 219 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
962diff
changeset | 220 | (* i(j ==> i(Suc(j) *) | 
| 923 | 221 | val less_SucI = lessI RSN (2, less_trans); | 
| 222 | ||
| 223 | goal Nat.thy "0 < Suc(n)"; | |
| 224 | by (nat_ind_tac "n" 1); | |
| 225 | by (rtac lessI 1); | |
| 226 | by (etac less_trans 1); | |
| 227 | by (rtac lessI 1); | |
| 228 | qed "zero_less_Suc"; | |
| 1301 | 229 | Addsimps [zero_less_Suc]; | 
| 923 | 230 | |
| 231 | (** Elimination properties **) | |
| 232 | ||
| 233 | val prems = goalw Nat.thy [less_def] "n<m ==> ~ m<(n::nat)"; | |
| 1552 | 234 | by (fast_tac (HOL_cs addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1); | 
| 923 | 235 | qed "less_not_sym"; | 
| 236 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
962diff
changeset | 237 | (* [| n(m; m(n |] ==> R *) | 
| 923 | 238 | bind_thm ("less_asym", (less_not_sym RS notE));
 | 
| 239 | ||
| 240 | goalw Nat.thy [less_def] "~ n<(n::nat)"; | |
| 241 | by (rtac notI 1); | |
| 242 | by (etac (wf_pred_nat RS wf_trancl RS wf_anti_refl) 1); | |
| 243 | qed "less_not_refl"; | |
| 244 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
962diff
changeset | 245 | (* n(n ==> R *) | 
| 923 | 246 | bind_thm ("less_anti_refl", (less_not_refl RS notE));
 | 
| 247 | ||
| 248 | goal Nat.thy "!!m. n<m ==> m ~= (n::nat)"; | |
| 1552 | 249 | by (fast_tac (HOL_cs addEs [less_anti_refl]) 1); | 
| 923 | 250 | qed "less_not_refl2"; | 
| 251 | ||
| 252 | ||
| 253 | val major::prems = goalw Nat.thy [less_def] | |
| 254 | "[| i<k; k=Suc(i) ==> P; !!j. [| i<j; k=Suc(j) |] ==> P \ | |
| 255 | \ |] ==> P"; | |
| 256 | by (rtac (major RS tranclE) 1); | |
| 1024 | 257 | by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE' | 
| 1465 | 258 | eresolve_tac (prems@[pred_natE, Pair_inject]))); | 
| 1024 | 259 | by (rtac refl 1); | 
| 923 | 260 | qed "lessE"; | 
| 261 | ||
| 262 | goal Nat.thy "~ n<0"; | |
| 263 | by (rtac notI 1); | |
| 264 | by (etac lessE 1); | |
| 265 | by (etac Zero_neq_Suc 1); | |
| 266 | by (etac Zero_neq_Suc 1); | |
| 267 | qed "not_less0"; | |
| 1301 | 268 | Addsimps [not_less0]; | 
| 923 | 269 | |
| 270 | (* n<0 ==> R *) | |
| 271 | bind_thm ("less_zeroE", (not_less0 RS notE));
 | |
| 272 | ||
| 273 | val [major,less,eq] = goal Nat.thy | |
| 274 | "[| m < Suc(n); m<n ==> P; m=n ==> P |] ==> P"; | |
| 275 | by (rtac (major RS lessE) 1); | |
| 276 | by (rtac eq 1); | |
| 277 | by (fast_tac (HOL_cs addSDs [Suc_inject]) 1); | |
| 278 | by (rtac less 1); | |
| 279 | by (fast_tac (HOL_cs addSDs [Suc_inject]) 1); | |
| 280 | qed "less_SucE"; | |
| 281 | ||
| 282 | goal Nat.thy "(m < Suc(n)) = (m < n | m = n)"; | |
| 283 | by (fast_tac (HOL_cs addSIs [lessI] | |
| 1465 | 284 | addEs [less_trans, less_SucE]) 1); | 
| 923 | 285 | qed "less_Suc_eq"; | 
| 286 | ||
| 1301 | 287 | val prems = goal Nat.thy "m<n ==> n ~= 0"; | 
| 1552 | 288 | by (res_inst_tac [("n","n")] natE 1);
 | 
| 289 | by (cut_facts_tac prems 1); | |
| 290 | by (ALLGOALS Asm_full_simp_tac); | |
| 1301 | 291 | qed "gr_implies_not0"; | 
| 292 | Addsimps [gr_implies_not0]; | |
| 923 | 293 | |
| 294 | (** Inductive (?) properties **) | |
| 295 | ||
| 296 | val [prem] = goal Nat.thy "Suc(m) < n ==> m<n"; | |
| 297 | by (rtac (prem RS rev_mp) 1); | |
| 298 | by (nat_ind_tac "n" 1); | |
| 299 | by (rtac impI 1); | |
| 300 | by (etac less_zeroE 1); | |
| 301 | by (fast_tac (HOL_cs addSIs [lessI RS less_SucI] | |
| 1465 | 302 | addSDs [Suc_inject] | 
| 303 | addEs [less_trans, lessE]) 1); | |
| 923 | 304 | qed "Suc_lessD"; | 
| 305 | ||
| 306 | val [major,minor] = goal Nat.thy | |
| 307 | "[| Suc(i)<k; !!j. [| i<j; k=Suc(j) |] ==> P \ | |
| 308 | \ |] ==> P"; | |
| 309 | by (rtac (major RS lessE) 1); | |
| 310 | by (etac (lessI RS minor) 1); | |
| 311 | by (etac (Suc_lessD RS minor) 1); | |
| 312 | by (assume_tac 1); | |
| 313 | qed "Suc_lessE"; | |
| 314 | ||
| 315 | val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m<n"; | |
| 316 | by (rtac (major RS lessE) 1); | |
| 317 | by (REPEAT (rtac lessI 1 | |
| 318 | ORELSE eresolve_tac [make_elim Suc_inject, ssubst, Suc_lessD] 1)); | |
| 319 | qed "Suc_less_SucD"; | |
| 320 | ||
| 321 | val prems = goal Nat.thy "m<n ==> Suc(m) < Suc(n)"; | |
| 322 | by (subgoal_tac "m<n --> Suc(m) < Suc(n)" 1); | |
| 323 | by (fast_tac (HOL_cs addIs prems) 1); | |
| 324 | by (nat_ind_tac "n" 1); | |
| 325 | by (rtac impI 1); | |
| 326 | by (etac less_zeroE 1); | |
| 327 | by (fast_tac (HOL_cs addSIs [lessI] | |
| 1465 | 328 | addSDs [Suc_inject] | 
| 329 | addEs [less_trans, lessE]) 1); | |
| 923 | 330 | qed "Suc_mono"; | 
| 331 | ||
| 332 | goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)"; | |
| 333 | by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]); | |
| 334 | qed "Suc_less_eq"; | |
| 1301 | 335 | Addsimps [Suc_less_eq]; | 
| 923 | 336 | |
| 337 | goal Nat.thy "~(Suc(n) < n)"; | |
| 1552 | 338 | by (fast_tac (HOL_cs addEs [Suc_lessD RS less_anti_refl]) 1); | 
| 923 | 339 | qed "not_Suc_n_less_n"; | 
| 1301 | 340 | Addsimps [not_Suc_n_less_n]; | 
| 341 | ||
| 342 | goal Nat.thy "!!i. i<j ==> j<k --> Suc i < k"; | |
| 1552 | 343 | by (nat_ind_tac "k" 1); | 
| 344 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [less_Suc_eq]))); | |
| 345 | by (fast_tac (HOL_cs addDs [Suc_lessD]) 1); | |
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1475diff
changeset | 346 | qed_spec_mp "less_trans_Suc"; | 
| 923 | 347 | |
| 348 | (*"Less than" is a linear ordering*) | |
| 349 | goal Nat.thy "m<n | m=n | n<(m::nat)"; | |
| 350 | by (nat_ind_tac "m" 1); | |
| 351 | by (nat_ind_tac "n" 1); | |
| 352 | by (rtac (refl RS disjI1 RS disjI2) 1); | |
| 353 | by (rtac (zero_less_Suc RS disjI1) 1); | |
| 354 | by (fast_tac (HOL_cs addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1); | |
| 355 | qed "less_linear"; | |
| 356 | ||
| 357 | (*Can be used with less_Suc_eq to get n=m | n<m *) | |
| 358 | goal Nat.thy "(~ m < n) = (n < Suc(m))"; | |
| 359 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | |
| 1552 | 360 | by (ALLGOALS Asm_simp_tac); | 
| 923 | 361 | qed "not_less_eq"; | 
| 362 | ||
| 363 | (*Complete induction, aka course-of-values induction*) | |
| 364 | val prems = goalw Nat.thy [less_def] | |
| 365 | "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |] ==> P(n)"; | |
| 366 | by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); | |
| 367 | by (eresolve_tac prems 1); | |
| 368 | qed "less_induct"; | |
| 369 | ||
| 370 | ||
| 371 | (*** Properties of <= ***) | |
| 372 | ||
| 373 | goalw Nat.thy [le_def] "0 <= n"; | |
| 374 | by (rtac not_less0 1); | |
| 375 | qed "le0"; | |
| 376 | ||
| 1301 | 377 | goalw Nat.thy [le_def] "~ Suc n <= n"; | 
| 1552 | 378 | by (Simp_tac 1); | 
| 1301 | 379 | qed "Suc_n_not_le_n"; | 
| 380 | ||
| 381 | goalw Nat.thy [le_def] "(i <= 0) = (i = 0)"; | |
| 1552 | 382 | by (nat_ind_tac "i" 1); | 
| 383 | by (ALLGOALS Asm_simp_tac); | |
| 1301 | 384 | qed "le_0"; | 
| 385 | ||
| 386 | Addsimps [less_not_refl, | |
| 387 | less_Suc_eq, le0, le_0, | |
| 388 | Suc_Suc_eq, Suc_n_not_le_n, | |
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 389 | n_not_Suc_n, Suc_n_not_n, | 
| 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 390 | nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; | 
| 923 | 391 | |
| 392 | (*Prevents simplification of f and g: much faster*) | |
| 393 | qed_goal "nat_case_weak_cong" Nat.thy | |
| 394 | "m=n ==> nat_case a f m = nat_case a f n" | |
| 395 | (fn [prem] => [rtac (prem RS arg_cong) 1]); | |
| 396 | ||
| 397 | qed_goal "nat_rec_weak_cong" Nat.thy | |
| 398 | "m=n ==> nat_rec m a f = nat_rec n a f" | |
| 399 | (fn [prem] => [rtac (prem RS arg_cong) 1]); | |
| 400 | ||
| 401 | val prems = goalw Nat.thy [le_def] "~(n<m) ==> m<=(n::nat)"; | |
| 402 | by (resolve_tac prems 1); | |
| 403 | qed "leI"; | |
| 404 | ||
| 405 | val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n<(m::nat))"; | |
| 406 | by (resolve_tac prems 1); | |
| 407 | qed "leD"; | |
| 408 | ||
| 409 | val leE = make_elim leD; | |
| 410 | ||
| 411 | goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)"; | |
| 412 | by (fast_tac HOL_cs 1); | |
| 413 | qed "not_leE"; | |
| 414 | ||
| 415 | goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n"; | |
| 1552 | 416 | by (Simp_tac 1); | 
| 923 | 417 | by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); | 
| 418 | qed "lessD"; | |
| 419 | ||
| 420 | goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n"; | |
| 1552 | 421 | by (Asm_full_simp_tac 1); | 
| 422 | by (fast_tac HOL_cs 1); | |
| 923 | 423 | qed "Suc_leD"; | 
| 424 | ||
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 425 | goalw Nat.thy [le_def] "!!m. m <= n ==> m <= Suc n"; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 426 | by (fast_tac (HOL_cs addDs [Suc_lessD]) 1); | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 427 | qed "le_SucI"; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 428 | Addsimps[le_SucI]; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 429 | |
| 923 | 430 | goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)"; | 
| 431 | by (fast_tac (HOL_cs addEs [less_asym]) 1); | |
| 432 | qed "less_imp_le"; | |
| 433 | ||
| 434 | goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)"; | |
| 435 | by (cut_facts_tac [less_linear] 1); | |
| 436 | by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); | |
| 437 | qed "le_imp_less_or_eq"; | |
| 438 | ||
| 439 | goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)"; | |
| 440 | by (cut_facts_tac [less_linear] 1); | |
| 441 | by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); | |
| 442 | by (flexflex_tac); | |
| 443 | qed "less_or_eq_imp_le"; | |
| 444 | ||
| 445 | goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)"; | |
| 446 | by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1)); | |
| 447 | qed "le_eq_less_or_eq"; | |
| 448 | ||
| 449 | goal Nat.thy "n <= (n::nat)"; | |
| 1552 | 450 | by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1); | 
| 923 | 451 | qed "le_refl"; | 
| 452 | ||
| 453 | val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)"; | |
| 454 | by (dtac le_imp_less_or_eq 1); | |
| 455 | by (fast_tac (HOL_cs addIs [less_trans]) 1); | |
| 456 | qed "le_less_trans"; | |
| 457 | ||
| 458 | goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)"; | |
| 459 | by (dtac le_imp_less_or_eq 1); | |
| 460 | by (fast_tac (HOL_cs addIs [less_trans]) 1); | |
| 461 | qed "less_le_trans"; | |
| 462 | ||
| 463 | goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)"; | |
| 464 | by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq, | |
| 465 | rtac less_or_eq_imp_le, fast_tac (HOL_cs addIs [less_trans])]); | |
| 466 | qed "le_trans"; | |
| 467 | ||
| 468 | val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)"; | |
| 469 | by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq, | |
| 470 | fast_tac (HOL_cs addEs [less_anti_refl,less_asym])]); | |
| 471 | qed "le_anti_sym"; | |
| 472 | ||
| 473 | goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)"; | |
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 474 | by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1); | 
| 923 | 475 | qed "Suc_le_mono"; | 
| 476 | ||
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1024diff
changeset | 477 | Addsimps [le_refl,Suc_le_mono]; | 
| 1531 | 478 | |
| 479 | ||
| 480 | (** LEAST -- the least number operator **) | |
| 481 | ||
| 482 | val [prem1,prem2] = goalw Nat.thy [Least_def] | |
| 483 | "[| P(k); !!x. x<k ==> ~P(x) |] ==> (LEAST x.P(x)) = k"; | |
| 484 | by (rtac select_equality 1); | |
| 485 | by (fast_tac (HOL_cs addSIs [prem1,prem2]) 1); | |
| 486 | by (cut_facts_tac [less_linear] 1); | |
| 487 | by (fast_tac (HOL_cs addSIs [prem1] addSDs [prem2]) 1); | |
| 488 | qed "Least_equality"; | |
| 489 | ||
| 490 | val [prem] = goal Nat.thy "P(k) ==> P(LEAST x.P(x))"; | |
| 491 | by (rtac (prem RS rev_mp) 1); | |
| 492 | by (res_inst_tac [("n","k")] less_induct 1);
 | |
| 493 | by (rtac impI 1); | |
| 494 | by (rtac classical 1); | |
| 495 | by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
 | |
| 496 | by (assume_tac 1); | |
| 497 | by (assume_tac 2); | |
| 498 | by (fast_tac HOL_cs 1); | |
| 499 | qed "LeastI"; | |
| 500 | ||
| 501 | (*Proof is almost identical to the one above!*) | |
| 502 | val [prem] = goal Nat.thy "P(k) ==> (LEAST x.P(x)) <= k"; | |
| 503 | by (rtac (prem RS rev_mp) 1); | |
| 504 | by (res_inst_tac [("n","k")] less_induct 1);
 | |
| 505 | by (rtac impI 1); | |
| 506 | by (rtac classical 1); | |
| 507 | by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
 | |
| 508 | by (assume_tac 1); | |
| 509 | by (rtac le_refl 2); | |
| 510 | by (fast_tac (HOL_cs addIs [less_imp_le,le_trans]) 1); | |
| 511 | qed "Least_le"; | |
| 512 | ||
| 513 | val [prem] = goal Nat.thy "k < (LEAST x.P(x)) ==> ~P(k)"; | |
| 514 | by (rtac notI 1); | |
| 515 | by (etac (rewrite_rule [le_def] Least_le RS notE) 1); | |
| 516 | by (rtac prem 1); | |
| 517 | qed "not_less_Least"; |