nat.ML
changeset 0 7949f97df77a
child 2 befa4e9f7c90
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/nat.ML	Thu Sep 16 12:21:07 1993 +0200
@@ -0,0 +1,409 @@
+(*  Title: 	HOL/nat
+    ID:         $Id$
+    Author: 	Tobias Nipkow, Cambridge University Computer Laboratory
+    Copyright   1991  University of Cambridge
+
+For nat.thy.  Type nat is defined as a set (Nat) over the type ind.
+*)
+
+open Nat;
+
+goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
+by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
+val Nat_fun_mono = result();
+
+val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
+
+(* Zero is a natural number -- this also justifies the type definition*)
+goal Nat.thy "Zero_Rep: Nat";
+by (rtac (Nat_unfold RS ssubst) 1);
+by (rtac (singletonI RS UnI1) 1);
+val Zero_RepI = result();
+
+val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat";
+by (rtac (Nat_unfold RS ssubst) 1);
+by (rtac (imageI RS UnI2) 1);
+by (resolve_tac prems 1);
+val Suc_RepI = result();
+
+(*** Induction ***)
+
+val major::prems = goal Nat.thy
+    "[| i: Nat;  P(Zero_Rep);   \
+\       !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |]  ==> P(i)";
+by (rtac (major RS (Nat_def RS def_induct)) 1);
+by (rtac Nat_fun_mono 1);
+by (fast_tac (set_cs addIs prems) 1);
+val Nat_induct = result();
+
+val prems = goalw Nat.thy [Zero_def,Suc_def]
+    "[| P(0);   \
+\       !!k. P(k) ==> P(Suc(k)) |]  ==> P(n)";
+by (rtac (Rep_Nat_inverse RS subst) 1);   (*types force good instantiation*)
+by (rtac (Rep_Nat RS Nat_induct) 1);
+by (REPEAT (ares_tac prems 1
+     ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1));
+val nat_induct = result();
+
+(*Perform induction on n. *)
+fun nat_ind_tac a i = 
+    EVERY [res_inst_tac [("n",a)] nat_induct i,
+	   rename_last_tac a ["1"] (i+1)];
+
+(*A special form of induction for reasoning about m<n and m-n*)
+val prems = goal Nat.thy
+    "[| !!x. P(x,0);  \
+\       !!y. P(0,Suc(y));  \
+\       !!x y. [| P(x,y) |] ==> P(Suc(x),Suc(y))  \
+\    |] ==> P(m,n)";
+by (res_inst_tac [("x","m")] spec 1);
+by (nat_ind_tac "n" 1);
+by (rtac allI 2);
+by (nat_ind_tac "x" 2);
+by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
+val diff_induct = result();
+
+(*Case analysis on the natural numbers*)
+val prems = goal Nat.thy 
+    "[| n=0 ==> P;  !!x. n = Suc(x) ==> P |] ==> P";
+by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1);
+by (fast_tac (HOL_cs addSEs prems) 1);
+by (nat_ind_tac "n" 1);
+by (rtac (refl RS disjI1) 1);
+by (fast_tac HOL_cs 1);
+val natE = result();
+
+(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
+
+(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat),
+  since we assume the isomorphism equations will one day be given by Isabelle*)
+
+goal Nat.thy "inj(Rep_Nat)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_Nat_inverse 1);
+val inj_Rep_Nat = result();
+
+goal Nat.thy "inj_onto(Abs_Nat,Nat)";
+by (rtac inj_onto_inverseI 1);
+by (etac Abs_Nat_inverse 1);
+val inj_onto_Abs_Nat = result();
+
+(*** Distinctness of constructors ***)
+
+goalw Nat.thy [Zero_def,Suc_def] "~ Suc(m)=0";
+by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1);
+by (rtac Suc_Rep_not_Zero_Rep 1);
+by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
+val Suc_not_Zero = result();
+
+val Zero_not_Suc = standard (Suc_not_Zero RS not_sym);
+
+val Suc_neq_Zero = standard (Suc_not_Zero RS notE);
+val Zero_neq_Suc = sym RS Suc_neq_Zero;
+
+(** Injectiveness of Suc **)
+
+goalw Nat.thy [Suc_def] "inj(Suc)";
+by (rtac injI 1);
+by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1);
+by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
+by (dtac (inj_Suc_Rep RS injD) 1);
+by (etac (inj_Rep_Nat RS injD) 1);
+val inj_Suc = result();
+
+val Suc_inject = inj_Suc RS injD;;
+
+goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)";
+by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); 
+val Suc_Suc_eq = result();
+
+goal Nat.thy "~ n=Suc(n)";
+by (nat_ind_tac "n" 1);
+by (ALLGOALS(asm_simp_tac (HOL_ss addsimps [Zero_not_Suc,Suc_Suc_eq])));
+val n_not_Suc_n = result();
+
+(*** nat_case -- the selection operator for nat ***)
+
+goalw Nat.thy [nat_case_def] "nat_case(0, a, f) = a";
+by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1);
+val nat_case_0 = result();
+
+goalw Nat.thy [nat_case_def] "nat_case(Suc(k), a, f) = f(k)";
+by (fast_tac (set_cs addIs [select_equality] 
+	               addEs [make_elim Suc_inject, Suc_neq_Zero]) 1);
+val nat_case_Suc = result();
+
+(** Introduction rules for 'pred_nat' **)
+
+goalw Nat.thy [pred_nat_def] "<n, Suc(n)> : pred_nat";
+by (fast_tac set_cs 1);
+val pred_natI = result();
+
+val major::prems = goalw Nat.thy [pred_nat_def]
+    "[| p : pred_nat;  !!x n. [| p = <n, Suc(n)> |] ==> R \
+\    |] ==> R";
+by (rtac (major RS CollectE) 1);
+by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1));
+val pred_natE = result();
+
+goalw Nat.thy [wf_def] "wf(pred_nat)";
+by (strip_tac 1);
+by (nat_ind_tac "x" 1);
+by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, 
+			     make_elim Suc_inject]) 2);
+by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, Zero_neq_Suc]) 1);
+val wf_pred_nat = result();
+
+
+(*** nat_rec -- by wf recursion on pred_nat ***)
+
+val nat_rec_unfold = standard (wf_pred_nat RS (nat_rec_def RS def_wfrec));
+
+(** conversion rules **)
+
+goal Nat.thy "nat_rec(0,c,h) = c";
+by (rtac (nat_rec_unfold RS trans) 1);
+by (rtac nat_case_0 1);
+val nat_rec_0 = result();
+
+goal Nat.thy "nat_rec(Suc(n), c, h) = h(n, nat_rec(n,c,h))";
+by (rtac (nat_rec_unfold RS trans) 1);
+by (rtac (nat_case_Suc RS trans) 1);
+by(simp_tac (HOL_ss addsimps [pred_natI,cut_apply]) 1);
+val nat_rec_Suc = result();
+
+(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
+val [rew] = goal Nat.thy
+    "[| !!n. f(n) == nat_rec(n,c,h) |] ==> f(0) = c";
+by (rewtac rew);
+by (rtac nat_rec_0 1);
+val def_nat_rec_0 = result();
+
+val [rew] = goal Nat.thy
+    "[| !!n. f(n) == nat_rec(n,c,h) |] ==> f(Suc(n)) = h(n,f(n))";
+by (rewtac rew);
+by (rtac nat_rec_Suc 1);
+val def_nat_rec_Suc = result();
+
+fun nat_recs def =
+      [standard (def RS def_nat_rec_0),
+       standard (def RS def_nat_rec_Suc)];
+
+
+(*** Basic properties of "less than" ***)
+
+(** Introduction properties **)
+
+val prems = goalw Nat.thy [less_def] "[| i<j;  j<k |] ==> i<k::nat";
+by (rtac (trans_trancl RS transD) 1);
+by (resolve_tac prems 1);
+by (resolve_tac prems 1);
+val less_trans = result();
+
+goalw Nat.thy [less_def] "n < Suc(n)";
+by (rtac (pred_natI RS r_into_trancl) 1);
+val lessI = result();
+
+(* i<j ==> i<Suc(j) *)
+val less_SucI = lessI RSN (2, less_trans);
+
+goal Nat.thy "0 < Suc(n)";
+by (nat_ind_tac "n" 1);
+by (rtac lessI 1);
+by (etac less_trans 1);
+by (rtac lessI 1);
+val zero_less_Suc = result();
+
+(** Elimination properties **)
+
+goalw Nat.thy [less_def] "n<m --> ~ m<n::nat";
+by (rtac (wf_pred_nat RS wf_trancl RS wf_anti_sym RS notI RS impI) 1);
+by (assume_tac 1);
+by (assume_tac 1);
+val less_not_sym = result();
+
+(* [| n<m; m<n |] ==> R *)
+val less_anti_sym = standard (less_not_sym RS mp RS notE);
+
+
+goalw Nat.thy [less_def] "~ n<n::nat";
+by (rtac notI 1);
+by (etac (wf_pred_nat RS wf_trancl RS wf_anti_refl) 1);
+val less_not_refl = result();
+
+(* n<n ==> R *)
+val less_anti_refl = standard (less_not_refl RS notE);
+
+
+val major::prems = goalw Nat.thy [less_def]
+    "[| i<k;  k=Suc(i) ==> P;  !!j. [| i<j;  k=Suc(j) |] ==> P \
+\    |] ==> P";
+by (rtac (major RS tranclE) 1);
+by (fast_tac (HOL_cs addSEs (prems@[pred_natE, Pair_inject])) 1);
+by (fast_tac (HOL_cs addSEs (prems@[pred_natE, Pair_inject])) 1);
+val lessE = result();
+
+goal Nat.thy "~ n<0";
+by (rtac notI 1);
+by (etac lessE 1);
+by (etac Zero_neq_Suc 1);
+by (etac Zero_neq_Suc 1);
+val not_less0 = result();
+
+(* n<0 ==> R *)
+val less_zeroE = standard (not_less0 RS notE);
+
+val [major,less,eq] = goal Nat.thy
+    "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
+by (rtac (major RS lessE) 1);
+by (rtac eq 1);
+by (fast_tac (HOL_cs addSDs [Suc_inject]) 1);
+by (rtac less 1);
+by (fast_tac (HOL_cs addSDs [Suc_inject]) 1);
+val less_SucE = result();
+
+goal Nat.thy "(m < Suc(n)) = (m < n | m = n)";
+by (fast_tac (HOL_cs addSIs [lessI]
+		     addEs  [less_trans, less_SucE]) 1);
+val less_Suc_eq = result();
+
+
+(** Inductive (?) properties **)
+
+val [prem] = goal Nat.thy "Suc(m) < n ==> m<n";
+by (rtac (prem RS rev_mp) 1);
+by (nat_ind_tac "n" 1);
+by (rtac impI 1);
+by (etac less_zeroE 1);
+by (fast_tac (HOL_cs addSIs [lessI RS less_SucI]
+	 	     addSDs [Suc_inject]
+		     addEs  [less_trans, lessE]) 1);
+val Suc_lessD = result();
+
+val [major,minor] = goal Nat.thy 
+    "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
+\    |] ==> P";
+by (rtac (major RS lessE) 1);
+by (etac (lessI RS minor) 1);
+by (etac (Suc_lessD RS minor) 1);
+by (assume_tac 1);
+val Suc_lessE = result();
+
+val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m<n";
+by (rtac (major RS lessE) 1);
+by (REPEAT (rtac lessI 1
+     ORELSE eresolve_tac [make_elim Suc_inject, ssubst, Suc_lessD] 1));
+val Suc_less_SucD = result();
+
+val prems = goal Nat.thy "m<n ==> Suc(m) < Suc(n)";
+by (subgoal_tac "m<n --> Suc(m) < Suc(n)" 1);
+by (fast_tac (HOL_cs addIs prems) 1);
+by (nat_ind_tac "n" 1);
+by (rtac impI 1);
+by (etac less_zeroE 1);
+by (fast_tac (HOL_cs addSIs [lessI]
+	 	     addSDs [Suc_inject]
+		     addEs  [less_trans, lessE]) 1);
+val Suc_mono = result();
+
+goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)";
+by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
+val Suc_less_eq = result();
+
+val [major] = goal Nat.thy "Suc(n)<n ==> P";
+by (rtac (major RS Suc_lessD RS less_anti_refl) 1);
+val not_Suc_n_less_n = result();
+
+(*"Less than" is a linear ordering*)
+goal Nat.thy "m<n | m=n | n<m::nat";
+by (nat_ind_tac "m" 1);
+by (nat_ind_tac "n" 1);
+by (rtac (refl RS disjI1 RS disjI2) 1);
+by (rtac (zero_less_Suc RS disjI1) 1);
+by (fast_tac (HOL_cs addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1);
+val less_linear = result();
+
+(*Can be used with less_Suc_eq to get n=m | n<m *)
+goal Nat.thy "(~ m < n) = (n < Suc(m))";
+by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
+by(ALLGOALS(asm_simp_tac (HOL_ss addsimps
+                          [not_less0,zero_less_Suc,Suc_less_eq])));
+val not_less_eq = result();
+
+(*Complete induction, aka course-of-values induction*)
+val prems = goalw Nat.thy [less_def]
+    "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
+by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
+by (eresolve_tac prems 1);
+val less_induct = result();
+
+
+(*** Properties of <= ***)
+
+goalw Nat.thy [le_def] "0 <= n";
+by (rtac not_less0 1);
+val le0 = result();
+
+val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI, 
+	     Suc_less_eq, less_Suc_eq, le0, 
+	     Suc_not_Zero, Zero_not_Suc, Suc_Suc_eq, 
+	     nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc];
+
+val nat_ss = pair_ss  addsimps  nat_simps;
+
+val prems = goalw Nat.thy [le_def] "~(n<m) ==> m<=n::nat";
+by (resolve_tac prems 1);
+val leI = result();
+
+val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n<m::nat)";
+by (resolve_tac prems 1);
+val leD = result();
+
+val leE = make_elim leD;
+
+goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<m::nat";
+by (fast_tac HOL_cs 1);
+val not_leE = result();
+
+goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n";
+by(simp_tac (HOL_ss addsimps [less_Suc_eq]) 1);
+by (fast_tac (HOL_cs addEs [less_anti_refl,less_anti_sym]) 1);
+val lessD = result();
+
+goalw Nat.thy [le_def] "!!m. m < n ==> m <= n::nat";
+by (fast_tac (HOL_cs addEs [less_anti_sym]) 1);
+val less_imp_le = result();
+
+goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=n::nat";
+by (cut_facts_tac [less_linear] 1);
+by (fast_tac (HOL_cs addEs [less_anti_refl,less_anti_sym]) 1);
+val le_imp_less_or_eq = result();
+
+goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <= n::nat";
+by (cut_facts_tac [less_linear] 1);
+by (fast_tac (HOL_cs addEs [less_anti_refl,less_anti_sym]) 1);
+by (flexflex_tac);
+val less_or_eq_imp_le = result();
+
+goal Nat.thy "(x <= y::nat) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
+val le_eq_less_or_eq = result();
+
+goal Nat.thy "n <= n::nat";
+by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1);
+val le_refl = result();
+
+val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < k::nat";
+by (dtac le_imp_less_or_eq 1);
+by (fast_tac (HOL_cs addIs [less_trans]) 1);
+val le_less_trans = result();
+
+goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= k::nat";
+by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
+          rtac less_or_eq_imp_le, fast_tac (HOL_cs addIs [less_trans])]);
+val le_trans = result();
+
+val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = n::nat";
+by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
+          fast_tac (HOL_cs addEs [less_anti_refl,less_anti_sym])]);
+val le_anti_sym = result();