clasohm@0: (* Title: ZF/nat.ML clasohm@0: ID: $Id$ clasohm@0: Author: Lawrence C Paulson, Cambridge University Computer Laboratory clasohm@0: Copyright 1992 University of Cambridge clasohm@0: clasohm@0: For nat.thy. Natural numbers in Zermelo-Fraenkel Set Theory clasohm@0: *) clasohm@0: clasohm@0: open Nat; clasohm@0: clasohm@0: goal Nat.thy "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"; clasohm@0: by (rtac bnd_monoI 1); clasohm@0: by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2)); clasohm@0: by (cut_facts_tac [infinity] 1); clasohm@0: by (fast_tac ZF_cs 1); clasohm@0: val nat_bnd_mono = result(); clasohm@0: clasohm@0: (* nat = {0} Un {succ(x). x:nat} *) clasohm@0: val nat_unfold = nat_bnd_mono RS (nat_def RS def_lfp_Tarski); clasohm@0: clasohm@0: (** Type checking of 0 and successor **) clasohm@0: clasohm@0: goal Nat.thy "0 : nat"; clasohm@0: by (rtac (nat_unfold RS ssubst) 1); clasohm@0: by (rtac (singletonI RS UnI1) 1); clasohm@0: val nat_0I = result(); clasohm@0: clasohm@0: val prems = goal Nat.thy "n : nat ==> succ(n) : nat"; clasohm@0: by (rtac (nat_unfold RS ssubst) 1); clasohm@0: by (rtac (RepFunI RS UnI2) 1); clasohm@0: by (resolve_tac prems 1); clasohm@0: val nat_succI = result(); clasohm@0: clasohm@0: goalw Nat.thy [one_def] "1 : nat"; clasohm@0: by (rtac (nat_0I RS nat_succI) 1); clasohm@0: val nat_1I = result(); clasohm@0: clasohm@0: goal Nat.thy "bool <= nat"; clasohm@0: by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 ORELSE etac boolE 1)); clasohm@0: val bool_subset_nat = result(); clasohm@0: clasohm@0: val bool_into_nat = bool_subset_nat RS subsetD; clasohm@0: clasohm@0: clasohm@0: (** Injectivity properties and induction **) clasohm@0: clasohm@0: (*Mathematical induction*) clasohm@0: val major::prems = goal Nat.thy clasohm@0: "[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"; clasohm@0: by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1); clasohm@0: by (fast_tac (ZF_cs addIs prems) 1); clasohm@0: val nat_induct = result(); clasohm@0: clasohm@0: (*Perform induction on n, then prove the n:nat subgoal using prems. *) clasohm@0: fun nat_ind_tac a prems i = clasohm@0: EVERY [res_inst_tac [("n",a)] nat_induct i, clasohm@0: rename_last_tac a ["1"] (i+2), clasohm@0: ares_tac prems i]; clasohm@0: clasohm@0: val major::prems = goal Nat.thy clasohm@0: "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"; clasohm@0: br (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1; clasohm@0: by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1 clasohm@0: ORELSE ares_tac prems 1)); clasohm@0: val natE = result(); clasohm@0: clasohm@0: val prems = goal Nat.thy "n: nat ==> Ord(n)"; clasohm@0: by (nat_ind_tac "n" prems 1); clasohm@0: by (REPEAT (ares_tac [Ord_0, Ord_succ] 1)); clasohm@0: val naturals_are_ordinals = result(); clasohm@0: clasohm@0: goal Nat.thy "!!n. n: nat ==> n=0 | 0:n"; clasohm@0: by (etac nat_induct 1); clasohm@0: by (fast_tac ZF_cs 1); clasohm@0: by (fast_tac (ZF_cs addIs [naturals_are_ordinals RS Ord_0_mem_succ]) 1); clasohm@0: val natE0 = result(); clasohm@0: clasohm@0: goal Nat.thy "Ord(nat)"; clasohm@0: by (rtac OrdI 1); clasohm@0: by (etac (naturals_are_ordinals RS Ord_is_Transset) 2); clasohm@0: by (rewtac Transset_def); clasohm@0: by (rtac ballI 1); clasohm@0: by (etac nat_induct 1); clasohm@0: by (REPEAT (ares_tac [empty_subsetI,succ_subsetI] 1)); clasohm@0: val Ord_nat = result(); clasohm@0: clasohm@0: (** Variations on mathematical induction **) clasohm@0: clasohm@0: (*complete induction*) clasohm@0: val complete_induct = Ord_nat RSN (2, Ord_induct); clasohm@0: clasohm@0: val prems = goal Nat.thy clasohm@0: "[| m: nat; n: nat; \ clasohm@0: \ !!x. [| x: nat; m<=x; P(x) |] ==> P(succ(x)) \ clasohm@0: \ |] ==> m <= n --> P(m) --> P(n)"; clasohm@0: by (nat_ind_tac "n" prems 1); clasohm@0: by (ALLGOALS clasohm@0: (ASM_SIMP_TAC clasohm@0: (ZF_ss addrews (prems@distrib_rews@[subset_empty_iff, subset_succ_iff, clasohm@0: Ord_nat RS Ord_in_Ord])))); clasohm@0: val nat_induct_from_lemma = result(); clasohm@0: clasohm@0: (*Induction starting from m rather than 0*) clasohm@0: val prems = goal Nat.thy clasohm@0: "[| m <= n; m: nat; n: nat; \ clasohm@0: \ P(m); \ clasohm@0: \ !!x. [| x: nat; m<=x; P(x) |] ==> P(succ(x)) \ clasohm@0: \ |] ==> P(n)"; clasohm@0: by (rtac (nat_induct_from_lemma RS mp RS mp) 1); clasohm@0: by (REPEAT (ares_tac prems 1)); clasohm@0: val nat_induct_from = result(); clasohm@0: clasohm@0: (*Induction suitable for subtraction and less-than*) clasohm@0: val prems = goal Nat.thy clasohm@0: "[| m: nat; n: nat; \ clasohm@0: \ !!x. [| x: nat |] ==> P(x,0); \ clasohm@0: \ !!y. [| y: nat |] ==> P(0,succ(y)); \ clasohm@0: \ !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) \ clasohm@0: \ |] ==> P(m,n)"; clasohm@0: by (res_inst_tac [("x","m")] bspec 1); clasohm@0: by (resolve_tac prems 2); clasohm@0: by (nat_ind_tac "n" prems 1); clasohm@0: by (rtac ballI 2); clasohm@0: by (nat_ind_tac "x" [] 2); clasohm@0: by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1)); clasohm@0: val diff_induct = result(); clasohm@0: clasohm@0: (** nat_case **) clasohm@0: clasohm@0: goalw Nat.thy [nat_case_def] "nat_case(0,a,b) = a"; clasohm@0: by (fast_tac (ZF_cs addIs [the_equality]) 1); clasohm@0: val nat_case_0 = result(); clasohm@0: clasohm@0: goalw Nat.thy [nat_case_def] "nat_case(succ(m),a,b) = b(m)"; clasohm@0: by (fast_tac (ZF_cs addIs [the_equality]) 1); clasohm@0: val nat_case_succ = result(); clasohm@0: clasohm@0: val major::prems = goal Nat.thy clasohm@0: "[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) \ clasohm@0: \ |] ==> nat_case(n,a,b) : C(n)"; clasohm@0: by (rtac (major RS nat_induct) 1); clasohm@0: by (REPEAT (resolve_tac [nat_case_0 RS ssubst, clasohm@0: nat_case_succ RS ssubst] 1 clasohm@0: THEN resolve_tac prems 1)); clasohm@0: by (assume_tac 1); clasohm@0: val nat_case_type = result(); clasohm@0: clasohm@0: val prems = goalw Nat.thy [nat_case_def] clasohm@0: "[| n=n'; a=a'; !!m z. b(m)=b'(m) \ clasohm@0: \ |] ==> nat_case(n,a,b)=nat_case(n',a',b')"; clasohm@0: by (REPEAT (resolve_tac [the_cong,disj_cong,ex_cong] 1 clasohm@0: ORELSE EVERY1 (map rtac ((prems RL [ssubst]) @ [iff_refl])))); clasohm@0: val nat_case_cong = result(); clasohm@0: clasohm@0: clasohm@0: (** nat_rec -- used to define eclose and transrec, then obsolete **) clasohm@0: clasohm@0: val nat_rec_trans = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); clasohm@0: clasohm@0: goal Nat.thy "nat_rec(0,a,b) = a"; clasohm@0: by (rtac nat_rec_trans 1); clasohm@0: by (rtac nat_case_0 1); clasohm@0: val nat_rec_0 = result(); clasohm@0: clasohm@0: val [prem] = goal Nat.thy clasohm@0: "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; clasohm@0: val nat_rec_ss = ZF_ss clasohm@0: addcongs (mk_typed_congs Nat.thy [("b", "[i,i]=>i")]) clasohm@0: addrews [prem, nat_case_succ, nat_succI, Memrel_iff, clasohm@0: vimage_singleton_iff]; clasohm@0: by (rtac nat_rec_trans 1); clasohm@0: by (SIMP_TAC nat_rec_ss 1); clasohm@0: val nat_rec_succ = result(); clasohm@0: clasohm@0: (** The union of two natural numbers is a natural number -- their maximum **) clasohm@0: clasohm@0: (* [| ?i : nat; ?j : nat |] ==> ?i Un ?j : nat *) clasohm@0: val Un_nat_type = standard (Ord_nat RSN (3,Ord_member_UnI)); clasohm@0: clasohm@0: (* [| ?i : nat; ?j : nat |] ==> ?i Int ?j : nat *) clasohm@0: val Int_nat_type = standard (Ord_nat RSN (3,Ord_member_IntI)); clasohm@0: