diff -r f04b33ce250f -r a4dc62a46ee4 Nat.ML --- a/Nat.ML Tue Oct 24 14:59:17 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,437 +0,0 @@ -(* 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)); -qed "Nat_fun_mono"; - -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); -qed "Zero_RepI"; - -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); -qed "Suc_RepI"; - -(*** 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 ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1); -by (fast_tac (set_cs addIs prems) 1); -qed "Nat_induct"; - -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)); -qed "nat_induct"; - -(*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 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)); -qed "diff_induct"; - -(*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); -qed "natE"; - -(*** 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); -qed "inj_Rep_Nat"; - -goal Nat.thy "inj_onto(Abs_Nat,Nat)"; -by (rtac inj_onto_inverseI 1); -by (etac Abs_Nat_inverse 1); -qed "inj_onto_Abs_Nat"; - -(*** 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)); -qed "Suc_not_Zero"; - -bind_thm ("Zero_not_Suc", (Suc_not_Zero RS not_sym)); - -bind_thm ("Suc_neq_Zero", (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); -qed "inj_Suc"; - -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]); -qed "Suc_Suc_eq"; - -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]))); -qed "n_not_Suc_n"; - -val Suc_n_not_n = n_not_Suc_n RS not_sym; - -(*** nat_case -- the selection operator for nat ***) - -goalw Nat.thy [nat_case_def] "nat_case(a, f, 0) = a"; -by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1); -qed "nat_case_0"; - -goalw Nat.thy [nat_case_def] "nat_case(a, f, Suc(k)) = f(k)"; -by (fast_tac (set_cs addIs [select_equality] - addEs [make_elim Suc_inject, Suc_neq_Zero]) 1); -qed "nat_case_Suc"; - -(** Introduction rules for 'pred_nat' **) - -goalw Nat.thy [pred_nat_def] " : pred_nat"; -by (fast_tac set_cs 1); -qed "pred_natI"; - -val major::prems = goalw Nat.thy [pred_nat_def] - "[| p : pred_nat; !!x n. [| p = |] ==> R \ -\ |] ==> R"; -by (rtac (major RS CollectE) 1); -by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1)); -qed "pred_natE"; - -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); -qed "wf_pred_nat"; - - -(*** nat_rec -- by wf recursion on pred_nat ***) - -bind_thm ("nat_rec_unfold", (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 (simp_tac (HOL_ss addsimps [nat_case_0]) 1); -qed "nat_rec_0"; - -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 (simp_tac (HOL_ss addsimps [nat_case_Suc, pred_natI, cut_apply]) 1); -qed "nat_rec_Suc"; - -(*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); -qed "def_nat_rec_0"; - -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); -qed "def_nat_rec_Suc"; - -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 i<(k::nat)"; -by (rtac (trans_trancl RS transD) 1); -by (resolve_tac prems 1); -by (resolve_tac prems 1); -qed "less_trans"; - -goalw Nat.thy [less_def] "n < Suc(n)"; -by (rtac (pred_natI RS r_into_trancl) 1); -qed "lessI"; - -(* i i ~ m<(n::nat)"; -by(fast_tac (HOL_cs addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1); -qed "less_not_sym"; - -(* [| n R *) -bind_thm ("less_asym", (less_not_sym 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); -qed "less_not_refl"; - -(* n R *) -bind_thm ("less_anti_refl", (less_not_refl RS notE)); - -goal Nat.thy "!!m. n m ~= (n::nat)"; -by(fast_tac (HOL_cs addEs [less_anti_refl]) 1); -qed "less_not_refl2"; - - -val major::prems = goalw Nat.thy [less_def] - "[| i P; !!j. [| i P \ -\ |] ==> P"; -by (rtac (major RS tranclE) 1); -by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE' - eresolve_tac (prems@[pred_natE, Pair_inject]))); -by (rtac refl 1); -qed "lessE"; - -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); -qed "not_less0"; - -(* n<0 ==> R *) -bind_thm ("less_zeroE", (not_less0 RS notE)); - -val [major,less,eq] = goal Nat.thy - "[| m < Suc(n); m 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); -qed "less_SucE"; - -goal Nat.thy "(m < Suc(n)) = (m < n | m = n)"; -by (fast_tac (HOL_cs addSIs [lessI] - addEs [less_trans, less_SucE]) 1); -qed "less_Suc_eq"; - - -(** Inductive (?) properties **) - -val [prem] = goal Nat.thy "Suc(m) < n ==> m 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); -qed "Suc_lessE"; - -val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m Suc(m) < Suc(n)"; -by (subgoal_tac "m 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); -qed "Suc_mono"; - -goal Nat.thy "(Suc(m) < Suc(n)) = (m P(m) |] ==> P(n) |] ==> P(n)"; -by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); -by (eresolve_tac prems 1); -qed "less_induct"; - - -(*** Properties of <= ***) - -goalw Nat.thy [le_def] "0 <= n"; -by (rtac not_less0 1); -qed "le0"; - -val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI, - Suc_less_eq, less_Suc_eq, le0, not_Suc_n_less_n, - Suc_not_Zero, Zero_not_Suc, Suc_Suc_eq, - n_not_Suc_n, Suc_n_not_n, - nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; - -val nat_ss0 = sum_ss addsimps nat_simps; - -(*Prevents simplification of f and g: much faster*) -qed_goal "nat_case_weak_cong" Nat.thy - "m=n ==> nat_case(a,f,m) = nat_case(a,f,n)" - (fn [prem] => [rtac (prem RS arg_cong) 1]); - -qed_goal "nat_rec_weak_cong" Nat.thy - "m=n ==> nat_rec(m,a,f) = nat_rec(n,a,f)" - (fn [prem] => [rtac (prem RS arg_cong) 1]); - -val prems = goalw Nat.thy [le_def] "~(n m<=(n::nat)"; -by (resolve_tac prems 1); -qed "leI"; - -val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n<(m::nat))"; -by (resolve_tac prems 1); -qed "leD"; - -val leE = make_elim leD; - -goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)"; -by (fast_tac HOL_cs 1); -qed "not_leE"; - -goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n"; -by(simp_tac nat_ss0 1); -by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); -qed "lessD"; - -goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n"; -by(asm_full_simp_tac nat_ss0 1); -by(fast_tac HOL_cs 1); -qed "Suc_leD"; - -goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)"; -by (fast_tac (HOL_cs addEs [less_asym]) 1); -qed "less_imp_le"; - -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_asym]) 1); -qed "le_imp_less_or_eq"; - -goalw Nat.thy [le_def] "!!m. m m <=(n::nat)"; -by (cut_facts_tac [less_linear] 1); -by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); -by (flexflex_tac); -qed "less_or_eq_imp_le"; - -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)); -qed "le_eq_less_or_eq"; - -goal Nat.thy "n <= (n::nat)"; -by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1); -qed "le_refl"; - -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); -qed "le_less_trans"; - -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); -qed "less_le_trans"; - -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])]); -qed "le_trans"; - -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_asym])]); -qed "le_anti_sym"; - -goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)"; -by (simp_tac (nat_ss0 addsimps [le_eq_less_or_eq]) 1); -qed "Suc_le_mono"; - -val nat_ss = nat_ss0 addsimps [le_refl,Suc_le_mono];