src/HOL/Nat.ML
changeset 2608 450c9b682a92
parent 2441 decc46a5cdb5
child 3023 01364e2f30ad
equal deleted inserted replaced
2607:a224a2865e05 2608:450c9b682a92
     1 (*  Title:      HOL/Nat.ML
     1 (*  Title:      HOL/Nat.ML
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Tobias Nipkow, Cambridge University Computer Laboratory
     3     Author:     Tobias Nipkow
     4     Copyright   1991  University of Cambridge
     4     Copyright   1997 TU Muenchen
     5 
       
     6 For Nat.thy.  Type nat is defined as a set (Nat) over the type ind.
       
     7 *)
     5 *)
     8 
     6 
     9 open Nat;
     7 goal thy "min 0 n = 0";
       
     8 br min_leastL 1;
       
     9 by(trans_tac 1);
       
    10 qed "min_0L";
    10 
    11 
    11 goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
    12 goal thy "min n 0 = 0";
    12 by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
    13 br min_leastR 1;
    13 qed "Nat_fun_mono";
    14 by(trans_tac 1);
       
    15 qed "min_0R";
    14 
    16 
    15 val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
    17 goalw thy [min_def] "min (Suc m) (Suc n) = Suc(min m n)";
       
    18 by(split_tac [expand_if] 1);
       
    19 by(Simp_tac 1);
       
    20 qed "min_Suc_Suc";
    16 
    21 
    17 (* Zero is a natural number -- this also justifies the type definition*)
    22 Addsimps [min_0L,min_0R,min_Suc_Suc];
    18 goal Nat.thy "Zero_Rep: Nat";
       
    19 by (stac Nat_unfold 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 (stac Nat_unfold 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 (!claset 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,
       
    50            rename_last_tac a ["1"] (i+1)];
       
    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 (!claset addSEs prems) 1);
       
    70 by (nat_ind_tac "n" 1);
       
    71 by (rtac (refl RS disjI1) 1);
       
    72 by (Fast_tac 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 
       
   100 AddIffs [Suc_not_Zero,Zero_not_Suc];
       
   101 
       
   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 
       
   115 val Suc_inject = inj_Suc RS injD;
       
   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 AddIffs [Suc_Suc_eq];
       
   122 
       
   123 goal Nat.thy "n ~= Suc(n)";
       
   124 by (nat_ind_tac "n" 1);
       
   125 by (ALLGOALS Asm_simp_tac);
       
   126 qed "n_not_Suc_n";
       
   127 
       
   128 bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
       
   129 
       
   130 (*** nat_case -- the selection operator for nat ***)
       
   131 
       
   132 goalw Nat.thy [nat_case_def] "nat_case a f 0 = a";
       
   133 by (fast_tac (!claset addIs [select_equality]) 1);
       
   134 qed "nat_case_0";
       
   135 
       
   136 goalw Nat.thy [nat_case_def] "nat_case a f (Suc k) = f(k)";
       
   137 by (fast_tac (!claset addIs [select_equality]) 1);
       
   138 qed "nat_case_Suc";
       
   139 
       
   140 (** Introduction rules for 'pred_nat' **)
       
   141 
       
   142 goalw Nat.thy [pred_nat_def] "(n, Suc(n)) : pred_nat";
       
   143 by (Fast_tac 1);
       
   144 qed "pred_natI";
       
   145 
       
   146 val major::prems = goalw Nat.thy [pred_nat_def]
       
   147     "[| p : pred_nat;  !!x n. [| p = (n, Suc(n)) |] ==> R \
       
   148 \    |] ==> R";
       
   149 by (rtac (major RS CollectE) 1);
       
   150 by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1));
       
   151 qed "pred_natE";
       
   152 
       
   153 goalw Nat.thy [wf_def] "wf(pred_nat)";
       
   154 by (strip_tac 1);
       
   155 by (nat_ind_tac "x" 1);
       
   156 by (fast_tac (!claset addSEs [mp, pred_natE]) 2);
       
   157 by (fast_tac (!claset addSEs [mp, pred_natE]) 1);
       
   158 qed "wf_pred_nat";
       
   159 
       
   160 
       
   161 (*** nat_rec -- by wf recursion on pred_nat ***)
       
   162 
       
   163 (* The unrolling rule for nat_rec *)
       
   164 goal Nat.thy
       
   165    "(%n. nat_rec c d n) = 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  *---------------------------------------------------------------------------*)
       
   174 
       
   175 (** conversion rules **)
       
   176 
       
   177 goal Nat.thy "nat_rec c h 0 = c";
       
   178 by (rtac (nat_rec_unfold RS trans) 1);
       
   179 by (simp_tac (!simpset addsimps [nat_case_0]) 1);
       
   180 qed "nat_rec_0";
       
   181 
       
   182 goal Nat.thy "nat_rec c h (Suc n) = h n (nat_rec c h n)";
       
   183 by (rtac (nat_rec_unfold RS trans) 1);
       
   184 by (simp_tac (!simpset addsimps [nat_case_Suc, pred_natI, cut_apply]) 1);
       
   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 c h n |] ==> 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 c h n |] ==> 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";
       
   218 AddIffs [lessI];
       
   219 
       
   220 (* i<j ==> i<Suc(j) *)
       
   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";
       
   229 AddIffs [zero_less_Suc];
       
   230 
       
   231 (** Elimination properties **)
       
   232 
       
   233 val prems = goalw Nat.thy [less_def] "n<m ==> ~ m<(n::nat)";
       
   234 by (fast_tac (!claset addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1);
       
   235 qed "less_not_sym";
       
   236 
       
   237 (* [| n<m; m<n |] ==> R *)
       
   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_irrefl) 1);
       
   243 qed "less_not_refl";
       
   244 
       
   245 (* n<n ==> R *)
       
   246 bind_thm ("less_irrefl", (less_not_refl RS notE));
       
   247 
       
   248 goal Nat.thy "!!m. n<m ==> m ~= (n::nat)";
       
   249 by (fast_tac (!claset addEs [less_irrefl]) 1);
       
   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);
       
   257 by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE'
       
   258                   eresolve_tac (prems@[pred_natE, Pair_inject])));
       
   259 by (rtac refl 1);
       
   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";
       
   268 
       
   269 AddIffs [not_less0];
       
   270 
       
   271 (* n<0 ==> R *)
       
   272 bind_thm ("less_zeroE", not_less0 RS notE);
       
   273 
       
   274 val [major,less,eq] = goal Nat.thy
       
   275     "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
       
   276 by (rtac (major RS lessE) 1);
       
   277 by (rtac eq 1);
       
   278 by (Fast_tac 1);
       
   279 by (rtac less 1);
       
   280 by (Fast_tac 1);
       
   281 qed "less_SucE";
       
   282 
       
   283 goal Nat.thy "(m < Suc(n)) = (m < n | m = n)";
       
   284 by (fast_tac (!claset addSIs [lessI]
       
   285                       addEs  [less_trans, less_SucE]) 1);
       
   286 qed "less_Suc_eq";
       
   287 
       
   288 val prems = goal Nat.thy "m<n ==> n ~= 0";
       
   289 by (res_inst_tac [("n","n")] natE 1);
       
   290 by (cut_facts_tac prems 1);
       
   291 by (ALLGOALS Asm_full_simp_tac);
       
   292 qed "gr_implies_not0";
       
   293 Addsimps [gr_implies_not0];
       
   294 
       
   295 qed_goal "zero_less_eq" Nat.thy "0 < n = (n ~= 0)" (fn _ => [
       
   296         rtac iffI 1,
       
   297         etac gr_implies_not0 1,
       
   298         rtac natE 1,
       
   299         contr_tac 1,
       
   300         etac ssubst 1,
       
   301         rtac zero_less_Suc 1]);
       
   302 
       
   303 (** Inductive (?) properties **)
       
   304 
       
   305 val [prem] = goal Nat.thy "Suc(m) < n ==> m<n";
       
   306 by (rtac (prem RS rev_mp) 1);
       
   307 by (nat_ind_tac "n" 1);
       
   308 by (ALLGOALS (fast_tac (!claset addSIs [lessI RS less_SucI]
       
   309                                 addEs  [less_trans, lessE])));
       
   310 qed "Suc_lessD";
       
   311 
       
   312 val [major,minor] = goal Nat.thy 
       
   313     "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
       
   314 \    |] ==> P";
       
   315 by (rtac (major RS lessE) 1);
       
   316 by (etac (lessI RS minor) 1);
       
   317 by (etac (Suc_lessD RS minor) 1);
       
   318 by (assume_tac 1);
       
   319 qed "Suc_lessE";
       
   320 
       
   321 goal Nat.thy "!!m n. Suc(m) < Suc(n) ==> m<n";
       
   322 by (fast_tac (!claset addEs [lessE, Suc_lessD] addIs [lessI]) 1);
       
   323 qed "Suc_less_SucD";
       
   324 
       
   325 goal Nat.thy "!!m n. m<n ==> Suc(m) < Suc(n)";
       
   326 by (etac rev_mp 1);
       
   327 by (nat_ind_tac "n" 1);
       
   328 by (ALLGOALS (fast_tac (!claset addSIs [lessI]
       
   329                                 addEs  [less_trans, lessE])));
       
   330 qed "Suc_mono";
       
   331 
       
   332 
       
   333 goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)";
       
   334 by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
       
   335 qed "Suc_less_eq";
       
   336 Addsimps [Suc_less_eq];
       
   337 
       
   338 goal Nat.thy "~(Suc(n) < n)";
       
   339 by (fast_tac (!claset addEs [Suc_lessD RS less_irrefl]) 1);
       
   340 qed "not_Suc_n_less_n";
       
   341 Addsimps [not_Suc_n_less_n];
       
   342 
       
   343 goal Nat.thy "!!i. i<j ==> j<k --> Suc i < k";
       
   344 by (nat_ind_tac "k" 1);
       
   345 by (ALLGOALS (asm_simp_tac (!simpset)));
       
   346 by (asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
       
   347 by (fast_tac (!claset addDs [Suc_lessD]) 1);
       
   348 qed_spec_mp "less_trans_Suc";
       
   349 
       
   350 (*"Less than" is a linear ordering*)
       
   351 goal Nat.thy "m<n | m=n | n<(m::nat)";
       
   352 by (nat_ind_tac "m" 1);
       
   353 by (nat_ind_tac "n" 1);
       
   354 by (rtac (refl RS disjI1 RS disjI2) 1);
       
   355 by (rtac (zero_less_Suc RS disjI1) 1);
       
   356 by (fast_tac (!claset addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1);
       
   357 qed "less_linear";
       
   358 
       
   359 qed_goal "nat_less_cases" Nat.thy 
       
   360    "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m"
       
   361 ( fn prems =>
       
   362         [
       
   363         (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1),
       
   364         (etac disjE 2),
       
   365         (etac (hd (tl (tl prems))) 1),
       
   366         (etac (sym RS hd (tl prems)) 1),
       
   367         (etac (hd prems) 1)
       
   368         ]);
       
   369 
       
   370 (*Can be used with less_Suc_eq to get n=m | n<m *)
       
   371 goal Nat.thy "(~ m < n) = (n < Suc(m))";
       
   372 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
       
   373 by (ALLGOALS Asm_simp_tac);
       
   374 qed "not_less_eq";
       
   375 
       
   376 (*Complete induction, aka course-of-values induction*)
       
   377 val prems = goalw Nat.thy [less_def]
       
   378     "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
       
   379 by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
       
   380 by (eresolve_tac prems 1);
       
   381 qed "less_induct";
       
   382 
       
   383 qed_goal "nat_induct2" Nat.thy 
       
   384 "[| P 0; P 1; !!k. P k ==> P (Suc (Suc k)) |] ==> P n" (fn prems => [
       
   385 	cut_facts_tac prems 1,
       
   386 	rtac less_induct 1,
       
   387 	res_inst_tac [("n","n")] natE 1,
       
   388 	 hyp_subst_tac 1,
       
   389 	 atac 1,
       
   390 	hyp_subst_tac 1,
       
   391 	res_inst_tac [("n","x")] natE 1,
       
   392 	 hyp_subst_tac 1,
       
   393 	 atac 1,
       
   394 	hyp_subst_tac 1,
       
   395 	resolve_tac prems 1,
       
   396 	dtac spec 1,
       
   397 	etac mp 1,
       
   398 	rtac (lessI RS less_trans) 1,
       
   399 	rtac (lessI RS Suc_mono) 1]);
       
   400 
       
   401 (*** Properties of <= ***)
       
   402 
       
   403 goalw Nat.thy [le_def] "(m <= n) = (m < Suc n)";
       
   404 by (rtac not_less_eq 1);
       
   405 qed "le_eq_less_Suc";
       
   406 
       
   407 goalw Nat.thy [le_def] "0 <= n";
       
   408 by (rtac not_less0 1);
       
   409 qed "le0";
       
   410 
       
   411 goalw Nat.thy [le_def] "~ Suc n <= n";
       
   412 by (Simp_tac 1);
       
   413 qed "Suc_n_not_le_n";
       
   414 
       
   415 goalw Nat.thy [le_def] "(i <= 0) = (i = 0)";
       
   416 by (nat_ind_tac "i" 1);
       
   417 by (ALLGOALS Asm_simp_tac);
       
   418 qed "le_0_eq";
       
   419 
       
   420 Addsimps [less_not_refl,
       
   421           (*less_Suc_eq, makes simpset non-confluent*) le0, le_0_eq,
       
   422           Suc_n_not_le_n,
       
   423           n_not_Suc_n, Suc_n_not_n,
       
   424           nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc];
       
   425 
       
   426 (*
       
   427 goal Nat.thy "(Suc m < n | Suc m = n) = (m < n)";
       
   428 by (stac (less_Suc_eq RS sym) 1);
       
   429 by (rtac Suc_less_eq 1);
       
   430 qed "Suc_le_eq";
       
   431 
       
   432 this could make the simpset (with less_Suc_eq added again) more confluent,
       
   433 but less_Suc_eq makes additional problems with terms of the form 0 < Suc (...)
       
   434 *)
       
   435 
       
   436 (*Prevents simplification of f and g: much faster*)
       
   437 qed_goal "nat_case_weak_cong" Nat.thy
       
   438   "m=n ==> nat_case a f m = nat_case a f n"
       
   439   (fn [prem] => [rtac (prem RS arg_cong) 1]);
       
   440 
       
   441 qed_goal "nat_rec_weak_cong" Nat.thy
       
   442   "m=n ==> nat_rec a f m = nat_rec a f n"
       
   443   (fn [prem] => [rtac (prem RS arg_cong) 1]);
       
   444 
       
   445 qed_goal "expand_nat_case" Nat.thy
       
   446   "P(nat_case z s n) = ((n=0 --> P z) & (!m. n = Suc m --> P(s m)))"
       
   447   (fn _ => [nat_ind_tac "n" 1, ALLGOALS Asm_simp_tac]);
       
   448 
       
   449 val prems = goalw Nat.thy [le_def] "~n<m ==> m<=(n::nat)";
       
   450 by (resolve_tac prems 1);
       
   451 qed "leI";
       
   452 
       
   453 val prems = goalw Nat.thy [le_def] "m<=n ==> ~ n < (m::nat)";
       
   454 by (resolve_tac prems 1);
       
   455 qed "leD";
       
   456 
       
   457 val leE = make_elim leD;
       
   458 
       
   459 goal Nat.thy "(~n<m) = (m<=(n::nat))";
       
   460 by (fast_tac (!claset addIs [leI] addEs [leE]) 1);
       
   461 qed "not_less_iff_le";
       
   462 
       
   463 goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)";
       
   464 by (Fast_tac 1);
       
   465 qed "not_leE";
       
   466 
       
   467 goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n";
       
   468 by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
       
   469 by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
       
   470 qed "lessD";
       
   471 
       
   472 goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n";
       
   473 by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
       
   474 qed "Suc_leD";
       
   475 
       
   476 (* stronger version of Suc_leD *)
       
   477 goalw Nat.thy [le_def] 
       
   478         "!!m. Suc m <= n ==> m < n";
       
   479 by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
       
   480 by (cut_facts_tac [less_linear] 1);
       
   481 by (Fast_tac 1);
       
   482 qed "Suc_le_lessD";
       
   483 
       
   484 goal Nat.thy "(Suc m <= n) = (m < n)";
       
   485 by (fast_tac (!claset addIs [lessD, Suc_le_lessD]) 1);
       
   486 qed "Suc_le_eq";
       
   487 
       
   488 goalw Nat.thy [le_def] "!!m. m <= n ==> m <= Suc n";
       
   489 by (fast_tac (!claset addDs [Suc_lessD]) 1);
       
   490 qed "le_SucI";
       
   491 Addsimps[le_SucI];
       
   492 
       
   493 bind_thm ("le_Suc", not_Suc_n_less_n RS leI);
       
   494 
       
   495 goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)";
       
   496 by (fast_tac (!claset addEs [less_asym]) 1);
       
   497 qed "less_imp_le";
       
   498 
       
   499 goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)";
       
   500 by (cut_facts_tac [less_linear] 1);
       
   501 by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
       
   502 qed "le_imp_less_or_eq";
       
   503 
       
   504 goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)";
       
   505 by (cut_facts_tac [less_linear] 1);
       
   506 by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
       
   507 by (flexflex_tac);
       
   508 qed "less_or_eq_imp_le";
       
   509 
       
   510 goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)";
       
   511 by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
       
   512 qed "le_eq_less_or_eq";
       
   513 
       
   514 goal Nat.thy "n <= (n::nat)";
       
   515 by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
       
   516 qed "le_refl";
       
   517 
       
   518 val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)";
       
   519 by (dtac le_imp_less_or_eq 1);
       
   520 by (fast_tac (!claset addIs [less_trans]) 1);
       
   521 qed "le_less_trans";
       
   522 
       
   523 goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)";
       
   524 by (dtac le_imp_less_or_eq 1);
       
   525 by (fast_tac (!claset addIs [less_trans]) 1);
       
   526 qed "less_le_trans";
       
   527 
       
   528 goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)";
       
   529 by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
       
   530           rtac less_or_eq_imp_le, fast_tac (!claset addIs [less_trans])]);
       
   531 qed "le_trans";
       
   532 
       
   533 val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)";
       
   534 by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
       
   535           fast_tac (!claset addEs [less_irrefl,less_asym])]);
       
   536 qed "le_anti_sym";
       
   537 
       
   538 goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)";
       
   539 by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
       
   540 qed "Suc_le_mono";
       
   541 
       
   542 AddIffs [le_refl,Suc_le_mono];
       
   543 
       
   544 
       
   545 (** LEAST -- the least number operator **)
       
   546 
       
   547 val [prem1,prem2] = goalw Nat.thy [Least_def]
       
   548     "[| P(k);  !!x. x<k ==> ~P(x) |] ==> (LEAST x.P(x)) = k";
       
   549 by (rtac select_equality 1);
       
   550 by (fast_tac (!claset addSIs [prem1,prem2]) 1);
       
   551 by (cut_facts_tac [less_linear] 1);
       
   552 by (fast_tac (!claset addSIs [prem1] addSDs [prem2]) 1);
       
   553 qed "Least_equality";
       
   554 
       
   555 val [prem] = goal Nat.thy "P(k) ==> P(LEAST x.P(x))";
       
   556 by (rtac (prem RS rev_mp) 1);
       
   557 by (res_inst_tac [("n","k")] less_induct 1);
       
   558 by (rtac impI 1);
       
   559 by (rtac classical 1);
       
   560 by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
       
   561 by (assume_tac 1);
       
   562 by (assume_tac 2);
       
   563 by (Fast_tac 1);
       
   564 qed "LeastI";
       
   565 
       
   566 (*Proof is almost identical to the one above!*)
       
   567 val [prem] = goal Nat.thy "P(k) ==> (LEAST x.P(x)) <= k";
       
   568 by (rtac (prem RS rev_mp) 1);
       
   569 by (res_inst_tac [("n","k")] less_induct 1);
       
   570 by (rtac impI 1);
       
   571 by (rtac classical 1);
       
   572 by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
       
   573 by (assume_tac 1);
       
   574 by (rtac le_refl 2);
       
   575 by (fast_tac (!claset addIs [less_imp_le,le_trans]) 1);
       
   576 qed "Least_le";
       
   577 
       
   578 val [prem] = goal Nat.thy "k < (LEAST x.P(x)) ==> ~P(k)";
       
   579 by (rtac notI 1);
       
   580 by (etac (rewrite_rule [le_def] Least_le RS notE) 1);
       
   581 by (rtac prem 1);
       
   582 qed "not_less_Least";
       
   583 
       
   584 qed_goalw "Least_Suc" Nat.thy [Least_def]
       
   585  "!!P. [| ? n. P(Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))"
       
   586  (fn _ => [
       
   587         rtac select_equality 1,
       
   588         fold_goals_tac [Least_def],
       
   589         safe_tac (!claset addSEs [LeastI]),
       
   590         res_inst_tac [("n","j")] natE 1,
       
   591         Fast_tac 1,
       
   592         fast_tac (!claset addDs [Suc_less_SucD, not_less_Least]) 1,
       
   593         res_inst_tac [("n","k")] natE 1,
       
   594         Fast_tac 1,
       
   595         hyp_subst_tac 1,
       
   596         rewtac Least_def,
       
   597         rtac (select_equality RS arg_cong RS sym) 1,
       
   598         safe_tac (!claset),
       
   599         dtac Suc_mono 1,
       
   600         Fast_tac 1,
       
   601         cut_facts_tac [less_linear] 1,
       
   602         safe_tac (!claset),
       
   603         atac 2,
       
   604         Fast_tac 2,
       
   605         dtac Suc_mono 1,
       
   606         Fast_tac 1]);
       
   607 
       
   608 
       
   609 (*** Instantiation of transitivity prover ***)
       
   610 
       
   611 structure Less_Arith =
       
   612 struct
       
   613 val nat_leI = leI;
       
   614 val nat_leD = leD;
       
   615 val lessI = lessI;
       
   616 val zero_less_Suc = zero_less_Suc;
       
   617 val less_reflE = less_irrefl;
       
   618 val less_zeroE = less_zeroE;
       
   619 val less_incr = Suc_mono;
       
   620 val less_decr = Suc_less_SucD;
       
   621 val less_incr_rhs = Suc_mono RS Suc_lessD;
       
   622 val less_decr_lhs = Suc_lessD;
       
   623 val less_trans_Suc = less_trans_Suc;
       
   624 val leI = lessD RS (Suc_le_mono RS iffD1);
       
   625 val not_lessI = leI RS leD
       
   626 val not_leI = prove_goal Nat.thy "!!m::nat. n < m ==> ~ m <= n"
       
   627   (fn _ => [etac swap2 1, etac leD 1]);
       
   628 val eqI = prove_goal Nat.thy "!!m. [| m < Suc n; n < Suc m |] ==> m=n"
       
   629   (fn _ => [etac less_SucE 1,
       
   630             fast_tac (HOL_cs addSDs [Suc_less_SucD] addSEs [less_irrefl]
       
   631                              addDs [less_trans_Suc]) 1,
       
   632             atac 1]);
       
   633 val leD = le_eq_less_Suc RS iffD1;
       
   634 val not_lessD = nat_leI RS leD;
       
   635 val not_leD = not_leE
       
   636 val eqD1 = prove_goal Nat.thy  "!!n. m = n ==> m < Suc n"
       
   637  (fn _ => [etac subst 1, rtac lessI 1]);
       
   638 val eqD2 = sym RS eqD1;
       
   639 
       
   640 fun is_zero(t) =  t = Const("0",Type("nat",[]));
       
   641 
       
   642 fun nnb T = T = Type("fun",[Type("nat",[]),
       
   643                             Type("fun",[Type("nat",[]),
       
   644                                         Type("bool",[])])])
       
   645 
       
   646 fun decomp_Suc(Const("Suc",_)$t) = let val (a,i) = decomp_Suc t in (a,i+1) end
       
   647   | decomp_Suc t = (t,0);
       
   648 
       
   649 fun decomp2(rel,T,lhs,rhs) =
       
   650   if not(nnb T) then None else
       
   651   let val (x,i) = decomp_Suc lhs
       
   652       val (y,j) = decomp_Suc rhs
       
   653   in case rel of
       
   654        "op <"  => Some(x,i,"<",y,j)
       
   655      | "op <=" => Some(x,i,"<=",y,j)
       
   656      | "op ="  => Some(x,i,"=",y,j)
       
   657      | _       => None
       
   658   end;
       
   659 
       
   660 fun negate(Some(x,i,rel,y,j)) = Some(x,i,"~"^rel,y,j)
       
   661   | negate None = None;
       
   662 
       
   663 fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp2(rel,T,lhs,rhs)
       
   664   | decomp(_$(Const("not",_)$(Const(rel,T)$lhs$rhs))) =
       
   665       negate(decomp2(rel,T,lhs,rhs))
       
   666   | decomp _ = None
       
   667 
       
   668 end;
       
   669 
       
   670 structure Trans_Tac = Trans_Tac_Fun(Less_Arith);
       
   671 
       
   672 open Trans_Tac;
       
   673 
       
   674 (*** eliminates ~= in premises, which trans_tac cannot deal with ***)
       
   675 qed_goal "nat_neqE" Nat.thy
       
   676   "[| (m::nat) ~= n; m < n ==> P; n < m ==> P |] ==> P"
       
   677   (fn major::prems => [cut_facts_tac [less_linear] 1,
       
   678                        REPEAT(eresolve_tac ([disjE,major RS notE]@prems) 1)]);