| author | berghofe | 
| Fri, 01 Oct 1999 10:23:13 +0200 | |
| changeset 7672 | c092e67d12f8 | 
| parent 120 | 09287f26bfb8 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: ZF/nat.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1992 University of Cambridge | |
| 5 | ||
| 6 | For nat.thy. Natural numbers in Zermelo-Fraenkel Set Theory | |
| 7 | *) | |
| 8 | ||
| 9 | open Nat; | |
| 10 | ||
| 11 | goal Nat.thy "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})";
 | |
| 12 | by (rtac bnd_monoI 1); | |
| 13 | by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2)); | |
| 14 | by (cut_facts_tac [infinity] 1); | |
| 15 | by (fast_tac ZF_cs 1); | |
| 16 | val nat_bnd_mono = result(); | |
| 17 | ||
| 18 | (* nat = {0} Un {succ(x). x:nat} *)
 | |
| 19 | val nat_unfold = nat_bnd_mono RS (nat_def RS def_lfp_Tarski); | |
| 20 | ||
| 21 | (** Type checking of 0 and successor **) | |
| 22 | ||
| 23 | goal Nat.thy "0 : nat"; | |
| 24 | by (rtac (nat_unfold RS ssubst) 1); | |
| 25 | by (rtac (singletonI RS UnI1) 1); | |
| 26 | val nat_0I = result(); | |
| 27 | ||
| 28 | val prems = goal Nat.thy "n : nat ==> succ(n) : nat"; | |
| 29 | by (rtac (nat_unfold RS ssubst) 1); | |
| 30 | by (rtac (RepFunI RS UnI2) 1); | |
| 31 | by (resolve_tac prems 1); | |
| 32 | val nat_succI = result(); | |
| 33 | ||
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 34 | goal Nat.thy "1 : nat"; | 
| 0 | 35 | by (rtac (nat_0I RS nat_succI) 1); | 
| 36 | val nat_1I = result(); | |
| 37 | ||
| 38 | goal Nat.thy "bool <= nat"; | |
| 120 | 39 | by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 | 
| 40 | ORELSE eresolve_tac [boolE,ssubst] 1)); | |
| 0 | 41 | val bool_subset_nat = result(); | 
| 42 | ||
| 43 | val bool_into_nat = bool_subset_nat RS subsetD; | |
| 44 | ||
| 45 | ||
| 46 | (** Injectivity properties and induction **) | |
| 47 | ||
| 48 | (*Mathematical induction*) | |
| 49 | val major::prems = goal Nat.thy | |
| 50 | "[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"; | |
| 51 | by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1); | |
| 52 | by (fast_tac (ZF_cs addIs prems) 1); | |
| 53 | val nat_induct = result(); | |
| 54 | ||
| 55 | (*Perform induction on n, then prove the n:nat subgoal using prems. *) | |
| 56 | fun nat_ind_tac a prems i = | |
| 57 |     EVERY [res_inst_tac [("n",a)] nat_induct i,
 | |
| 58 | rename_last_tac a ["1"] (i+2), | |
| 59 | ares_tac prems i]; | |
| 60 | ||
| 61 | val major::prems = goal Nat.thy | |
| 62 | "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"; | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 63 | by (rtac (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1); | 
| 0 | 64 | by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1 | 
| 65 | ORELSE ares_tac prems 1)); | |
| 66 | val natE = result(); | |
| 67 | ||
| 68 | val prems = goal Nat.thy "n: nat ==> Ord(n)"; | |
| 69 | by (nat_ind_tac "n" prems 1); | |
| 70 | by (REPEAT (ares_tac [Ord_0, Ord_succ] 1)); | |
| 71 | val naturals_are_ordinals = result(); | |
| 72 | ||
| 30 | 73 | (* i: nat ==> 0 le i *) | 
| 74 | val nat_0_le = naturals_are_ordinals RS Ord_0_le; | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 75 | |
| 0 | 76 | goal Nat.thy "!!n. n: nat ==> n=0 | 0:n"; | 
| 77 | by (etac nat_induct 1); | |
| 78 | by (fast_tac ZF_cs 1); | |
| 30 | 79 | by (fast_tac (ZF_cs addIs [nat_0_le]) 1); | 
| 0 | 80 | val natE0 = result(); | 
| 81 | ||
| 82 | goal Nat.thy "Ord(nat)"; | |
| 83 | by (rtac OrdI 1); | |
| 84 | by (etac (naturals_are_ordinals RS Ord_is_Transset) 2); | |
| 85 | by (rewtac Transset_def); | |
| 86 | by (rtac ballI 1); | |
| 87 | by (etac nat_induct 1); | |
| 88 | by (REPEAT (ares_tac [empty_subsetI,succ_subsetI] 1)); | |
| 89 | val Ord_nat = result(); | |
| 90 | ||
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 91 | (* succ(i): nat ==> i: nat *) | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 92 | val succ_natD = [succI1, asm_rl, Ord_nat] MRS Ord_trans; | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 93 | |
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 94 | (* [| succ(i): k; k: nat |] ==> i: k *) | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 95 | val succ_in_naturalD = [succI1, asm_rl, naturals_are_ordinals] MRS Ord_trans; | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 96 | |
| 30 | 97 | goal Nat.thy "!!m n. [| m<n; n: nat |] ==> m: nat"; | 
| 98 | by (etac ltE 1); | |
| 99 | by (etac (Ord_nat RSN (3,Ord_trans)) 1); | |
| 100 | by (assume_tac 1); | |
| 101 | val lt_nat_in_nat = result(); | |
| 102 | ||
| 103 | ||
| 0 | 104 | (** Variations on mathematical induction **) | 
| 105 | ||
| 106 | (*complete induction*) | |
| 107 | val complete_induct = Ord_nat RSN (2, Ord_induct); | |
| 108 | ||
| 109 | val prems = goal Nat.thy | |
| 110 | "[| m: nat; n: nat; \ | |
| 30 | 111 | \ !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) \ | 
| 112 | \ |] ==> m le n --> P(m) --> P(n)"; | |
| 0 | 113 | by (nat_ind_tac "n" prems 1); | 
| 114 | by (ALLGOALS | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 115 | (asm_simp_tac | 
| 30 | 116 | (ZF_ss addsimps (prems@distrib_rews@[le0_iff, le_succ_iff])))); | 
| 0 | 117 | val nat_induct_from_lemma = result(); | 
| 118 | ||
| 119 | (*Induction starting from m rather than 0*) | |
| 120 | val prems = goal Nat.thy | |
| 30 | 121 | "[| m le n; m: nat; n: nat; \ | 
| 0 | 122 | \ P(m); \ | 
| 30 | 123 | \ !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) \ | 
| 0 | 124 | \ |] ==> P(n)"; | 
| 125 | by (rtac (nat_induct_from_lemma RS mp RS mp) 1); | |
| 126 | by (REPEAT (ares_tac prems 1)); | |
| 127 | val nat_induct_from = result(); | |
| 128 | ||
| 129 | (*Induction suitable for subtraction and less-than*) | |
| 130 | val prems = goal Nat.thy | |
| 131 | "[| m: nat; n: nat; \ | |
| 30 | 132 | \ !!x. x: nat ==> P(x,0); \ | 
| 133 | \ !!y. y: nat ==> P(0,succ(y)); \ | |
| 0 | 134 | \ !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) \ | 
| 135 | \ |] ==> P(m,n)"; | |
| 136 | by (res_inst_tac [("x","m")] bspec 1);
 | |
| 137 | by (resolve_tac prems 2); | |
| 138 | by (nat_ind_tac "n" prems 1); | |
| 139 | by (rtac ballI 2); | |
| 140 | by (nat_ind_tac "x" [] 2); | |
| 141 | by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1)); | |
| 142 | val diff_induct = result(); | |
| 143 | ||
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 144 | (** Induction principle analogous to trancl_induct **) | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 145 | |
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 146 | goal Nat.thy | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 147 | "!!m. m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> \ | 
| 30 | 148 | \ (ALL n:nat. m<n --> P(m,n))"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 149 | by (etac nat_induct 1); | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 150 | by (ALLGOALS | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 151 | (EVERY' [rtac (impI RS impI), rtac (nat_induct RS ballI), assume_tac, | 
| 30 | 152 | fast_tac lt_cs, fast_tac lt_cs])); | 
| 153 | val succ_lt_induct_lemma = result(); | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 154 | |
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 155 | val prems = goal Nat.thy | 
| 30 | 156 | "[| m<n; n: nat; \ | 
| 157 | \ P(m,succ(m)); \ | |
| 158 | \ !!x. [| x: nat; P(m,x) |] ==> P(m,succ(x)) \ | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 159 | \ |] ==> P(m,n)"; | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 160 | by (res_inst_tac [("P4","P")] 
 | 
| 30 | 161 | (succ_lt_induct_lemma RS mp RS mp RS bspec RS mp) 1); | 
| 162 | by (REPEAT (ares_tac (prems @ [ballI, impI, lt_nat_in_nat]) 1)); | |
| 163 | val succ_lt_induct = result(); | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 164 | |
| 0 | 165 | (** nat_case **) | 
| 166 | ||
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 167 | goalw Nat.thy [nat_case_def] "nat_case(a,b,0) = a"; | 
| 0 | 168 | by (fast_tac (ZF_cs addIs [the_equality]) 1); | 
| 169 | val nat_case_0 = result(); | |
| 170 | ||
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 171 | goalw Nat.thy [nat_case_def] "nat_case(a,b,succ(m)) = b(m)"; | 
| 0 | 172 | by (fast_tac (ZF_cs addIs [the_equality]) 1); | 
| 173 | val nat_case_succ = result(); | |
| 174 | ||
| 175 | val major::prems = goal Nat.thy | |
| 176 | "[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) \ | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 177 | \ |] ==> nat_case(a,b,n) : C(n)"; | 
| 0 | 178 | by (rtac (major RS nat_induct) 1); | 
| 30 | 179 | by (ALLGOALS | 
| 180 | (asm_simp_tac (ZF_ss addsimps (prems @ [nat_case_0, nat_case_succ])))); | |
| 0 | 181 | val nat_case_type = result(); | 
| 182 | ||
| 183 | ||
| 30 | 184 | (** nat_rec -- used to define eclose and transrec, then obsolete; | 
| 185 | rec, from arith.ML, has fewer typing conditions **) | |
| 0 | 186 | |
| 187 | val nat_rec_trans = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); | |
| 188 | ||
| 189 | goal Nat.thy "nat_rec(0,a,b) = a"; | |
| 190 | by (rtac nat_rec_trans 1); | |
| 191 | by (rtac nat_case_0 1); | |
| 192 | val nat_rec_0 = result(); | |
| 193 | ||
| 194 | val [prem] = goal Nat.thy | |
| 195 | "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; | |
| 196 | by (rtac nat_rec_trans 1); | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 197 | by (simp_tac (ZF_ss addsimps [prem, nat_case_succ, nat_succI, Memrel_iff, | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 198 | vimage_singleton_iff]) 1); | 
| 0 | 199 | val nat_rec_succ = result(); | 
| 200 | ||
| 201 | (** The union of two natural numbers is a natural number -- their maximum **) | |
| 202 | ||
| 30 | 203 | goal Nat.thy "!!i j. [| i: nat; j: nat |] ==> i Un j: nat"; | 
| 204 | by (rtac (Un_least_lt RS ltD) 1); | |
| 205 | by (REPEAT (ares_tac [ltI, Ord_nat] 1)); | |
| 206 | val Un_nat_type = result(); | |
| 0 | 207 | |
| 30 | 208 | goal Nat.thy "!!i j. [| i: nat; j: nat |] ==> i Int j: nat"; | 
| 209 | by (rtac (Int_greatest_lt RS ltD) 1); | |
| 210 | by (REPEAT (ares_tac [ltI, Ord_nat] 1)); | |
| 211 | val Int_nat_type = result(); |