| author | paulson | 
| Fri, 25 Sep 1998 13:57:01 +0200 | |
| changeset 5562 | 02261e6880d1 | 
| parent 5529 | 4a54acae6a15 | 
| child 6070 | 032babd0120b | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/nat.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 5505 | 6 | Natural numbers in Zermelo-Fraenkel Set Theory | 
| 0 | 7 | *) | 
| 8 | ||
| 9 | open Nat; | |
| 10 | ||
| 5067 | 11 | Goal "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})";
 | 
| 0 | 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); | |
| 4243 | 15 | by (Blast_tac 1); | 
| 760 | 16 | qed "nat_bnd_mono"; | 
| 0 | 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 | ||
| 5067 | 23 | Goal "0 : nat"; | 
| 2033 | 24 | by (stac nat_unfold 1); | 
| 0 | 25 | by (rtac (singletonI RS UnI1) 1); | 
| 760 | 26 | qed "nat_0I"; | 
| 0 | 27 | |
| 5137 | 28 | Goal "n : nat ==> succ(n) : nat"; | 
| 2033 | 29 | by (stac nat_unfold 1); | 
| 5137 | 30 | by (etac (RepFunI RS UnI2) 1); | 
| 760 | 31 | qed "nat_succI"; | 
| 0 | 32 | |
| 5067 | 33 | Goal "1 : nat"; | 
| 0 | 34 | by (rtac (nat_0I RS nat_succI) 1); | 
| 760 | 35 | qed "nat_1I"; | 
| 0 | 36 | |
| 5067 | 37 | Goal "2 : nat"; | 
| 829 | 38 | by (rtac (nat_1I RS nat_succI) 1); | 
| 39 | qed "nat_2I"; | |
| 40 | ||
| 5137 | 41 | Addsimps [nat_0I, nat_1I, nat_2I]; | 
| 42 | AddSIs [nat_0I, nat_1I, nat_2I, nat_succI]; | |
| 2469 | 43 | |
| 5067 | 44 | Goal "bool <= nat"; | 
| 120 | 45 | by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 | 
| 1461 | 46 | ORELSE eresolve_tac [boolE,ssubst] 1)); | 
| 760 | 47 | qed "bool_subset_nat"; | 
| 0 | 48 | |
| 49 | val bool_into_nat = bool_subset_nat RS subsetD; | |
| 50 | ||
| 51 | ||
| 52 | (** Injectivity properties and induction **) | |
| 53 | ||
| 54 | (*Mathematical induction*) | |
| 55 | val major::prems = goal Nat.thy | |
| 56 | "[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"; | |
| 57 | by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1); | |
| 4243 | 58 | by (blast_tac (claset() addIs prems) 1); | 
| 760 | 59 | qed "nat_induct"; | 
| 0 | 60 | |
| 61 | (*Perform induction on n, then prove the n:nat subgoal using prems. *) | |
| 62 | fun nat_ind_tac a prems i = | |
| 63 |     EVERY [res_inst_tac [("n",a)] nat_induct i,
 | |
| 1461 | 64 | rename_last_tac a ["1"] (i+2), | 
| 65 | ares_tac prems i]; | |
| 0 | 66 | |
| 67 | val major::prems = goal Nat.thy | |
| 68 | "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"; | |
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 69 | by (rtac (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1); | 
| 0 | 70 | by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1 | 
| 71 | ORELSE ares_tac prems 1)); | |
| 760 | 72 | qed "natE"; | 
| 0 | 73 | |
| 74 | val prems = goal Nat.thy "n: nat ==> Ord(n)"; | |
| 75 | by (nat_ind_tac "n" prems 1); | |
| 76 | by (REPEAT (ares_tac [Ord_0, Ord_succ] 1)); | |
| 760 | 77 | qed "nat_into_Ord"; | 
| 0 | 78 | |
| 1610 | 79 | (* i: nat ==> 0 le i; same thing as 0<succ(i) *) | 
| 80 | bind_thm ("nat_0_le", nat_into_Ord RS Ord_0_le);
 | |
| 435 | 81 | |
| 1610 | 82 | (* i: nat ==> i le i; same thing as i<succ(i) *) | 
| 83 | bind_thm ("nat_le_refl", nat_into_Ord RS le_refl);
 | |
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 84 | |
| 5067 | 85 | Goal "Ord(nat)"; | 
| 0 | 86 | by (rtac OrdI 1); | 
| 435 | 87 | by (etac (nat_into_Ord RS Ord_is_Transset) 2); | 
| 0 | 88 | by (rewtac Transset_def); | 
| 89 | by (rtac ballI 1); | |
| 90 | by (etac nat_induct 1); | |
| 91 | by (REPEAT (ares_tac [empty_subsetI,succ_subsetI] 1)); | |
| 760 | 92 | qed "Ord_nat"; | 
| 0 | 93 | |
| 5067 | 94 | Goalw [Limit_def] "Limit(nat)"; | 
| 5137 | 95 | by (safe_tac (claset() addSIs [ltI, Ord_nat])); | 
| 435 | 96 | by (etac ltD 1); | 
| 760 | 97 | qed "Limit_nat"; | 
| 435 | 98 | |
| 5137 | 99 | (* succ(i): nat ==> i: nat *) | 
| 100 | val succ_natD = [succI1, asm_rl, Ord_nat] MRS Ord_trans; | |
| 101 | AddSDs [succ_natD]; | |
| 102 | ||
| 103 | Goal "succ(n): nat <-> n: nat"; | |
| 104 | by (Blast_tac 1); | |
| 105 | qed "nat_succ_iff"; | |
| 106 | ||
| 107 | Addsimps [Ord_nat, Limit_nat, nat_succ_iff]; | |
| 2469 | 108 | AddSIs [Ord_nat, Limit_nat]; | 
| 109 | ||
| 5137 | 110 | Goal "Limit(i) ==> nat le i"; | 
| 1461 | 111 | by (rtac subset_imp_le 1); | 
| 484 | 112 | by (rtac subsetI 1); | 
| 1461 | 113 | by (etac nat_induct 1); | 
| 4091 | 114 | by (blast_tac (claset() addIs [Limit_has_succ RS ltD, ltI, Limit_is_Ord]) 2); | 
| 484 | 115 | by (REPEAT (ares_tac [Limit_has_0 RS ltD, | 
| 1461 | 116 | Ord_nat, Limit_is_Ord] 1)); | 
| 760 | 117 | qed "nat_le_Limit"; | 
| 484 | 118 | |
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 119 | (* [| succ(i): k; k: nat |] ==> i: k *) | 
| 435 | 120 | val succ_in_naturalD = [succI1, asm_rl, nat_into_Ord] MRS Ord_trans; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 121 | |
| 5137 | 122 | Goal "[| m<n; n: nat |] ==> m: nat"; | 
| 30 | 123 | by (etac ltE 1); | 
| 124 | by (etac (Ord_nat RSN (3,Ord_trans)) 1); | |
| 125 | by (assume_tac 1); | |
| 760 | 126 | qed "lt_nat_in_nat"; | 
| 30 | 127 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5147diff
changeset | 128 | Goal "[| m le n; n:nat |] ==> m:nat"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5147diff
changeset | 129 | by (blast_tac (claset() addSDs [lt_nat_in_nat]) 1); | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5147diff
changeset | 130 | qed "le_in_nat"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5147diff
changeset | 131 | |
| 30 | 132 | |
| 0 | 133 | (** Variations on mathematical induction **) | 
| 134 | ||
| 135 | (*complete induction*) | |
| 136 | val complete_induct = Ord_nat RSN (2, Ord_induct); | |
| 137 | ||
| 138 | val prems = goal Nat.thy | |
| 139 | "[| m: nat; n: nat; \ | |
| 30 | 140 | \ !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) \ | 
| 141 | \ |] ==> m le n --> P(m) --> P(n)"; | |
| 0 | 142 | by (nat_ind_tac "n" prems 1); | 
| 143 | by (ALLGOALS | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 144 | (asm_simp_tac | 
| 5529 | 145 | (simpset() addsimps prems @ distrib_simps @ [le0_iff, le_succ_iff]))); | 
| 760 | 146 | qed "nat_induct_from_lemma"; | 
| 0 | 147 | |
| 148 | (*Induction starting from m rather than 0*) | |
| 149 | val prems = goal Nat.thy | |
| 30 | 150 | "[| m le n; m: nat; n: nat; \ | 
| 0 | 151 | \ P(m); \ | 
| 30 | 152 | \ !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) \ | 
| 0 | 153 | \ |] ==> P(n)"; | 
| 154 | by (rtac (nat_induct_from_lemma RS mp RS mp) 1); | |
| 155 | by (REPEAT (ares_tac prems 1)); | |
| 760 | 156 | qed "nat_induct_from"; | 
| 0 | 157 | |
| 158 | (*Induction suitable for subtraction and less-than*) | |
| 159 | val prems = goal Nat.thy | |
| 160 | "[| m: nat; n: nat; \ | |
| 30 | 161 | \ !!x. x: nat ==> P(x,0); \ | 
| 162 | \ !!y. y: nat ==> P(0,succ(y)); \ | |
| 0 | 163 | \ !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) \ | 
| 164 | \ |] ==> P(m,n)"; | |
| 165 | by (res_inst_tac [("x","m")] bspec 1);
 | |
| 166 | by (resolve_tac prems 2); | |
| 167 | by (nat_ind_tac "n" prems 1); | |
| 168 | by (rtac ballI 2); | |
| 169 | by (nat_ind_tac "x" [] 2); | |
| 170 | by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1)); | |
| 760 | 171 | qed "diff_induct"; | 
| 0 | 172 | |
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 173 | (** Induction principle analogous to trancl_induct **) | 
| 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 174 | |
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 175 | Goal "m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> \ | 
| 30 | 176 | \ (ALL n:nat. m<n --> P(m,n))"; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 177 | by (etac nat_induct 1); | 
| 5479 | 178 | by (ALLGOALS | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 179 | (EVERY' [rtac (impI RS impI), rtac (nat_induct RS ballI), assume_tac, | 
| 5479 | 180 | blast_tac le_cs, blast_tac le_cs])); | 
| 760 | 181 | qed "succ_lt_induct_lemma"; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 182 | |
| 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 183 | val prems = goal Nat.thy | 
| 1461 | 184 | "[| m<n; n: nat; \ | 
| 185 | \ P(m,succ(m)); \ | |
| 186 | \ !!x. [| x: nat; P(m,x) |] ==> P(m,succ(x)) \ | |
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 187 | \ |] ==> P(m,n)"; | 
| 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 188 | by (res_inst_tac [("P4","P")] 
 | 
| 30 | 189 | (succ_lt_induct_lemma RS mp RS mp RS bspec RS mp) 1); | 
| 190 | by (REPEAT (ares_tac (prems @ [ballI, impI, lt_nat_in_nat]) 1)); | |
| 760 | 191 | qed "succ_lt_induct"; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 192 | |
| 0 | 193 | (** nat_case **) | 
| 194 | ||
| 5067 | 195 | Goalw [nat_case_def] "nat_case(a,b,0) = a"; | 
| 5505 | 196 | by (Blast_tac 1); | 
| 760 | 197 | qed "nat_case_0"; | 
| 0 | 198 | |
| 5067 | 199 | Goalw [nat_case_def] "nat_case(a,b,succ(m)) = b(m)"; | 
| 5505 | 200 | by (Blast_tac 1); | 
| 760 | 201 | qed "nat_case_succ"; | 
| 0 | 202 | |
| 2469 | 203 | Addsimps [nat_case_0, nat_case_succ]; | 
| 204 | ||
| 0 | 205 | val major::prems = goal Nat.thy | 
| 206 | "[| 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 | 207 | \ |] ==> nat_case(a,b,n) : C(n)"; | 
| 0 | 208 | by (rtac (major RS nat_induct) 1); | 
| 4091 | 209 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 210 | qed "nat_case_type"; | 
| 0 | 211 | |
| 212 | ||
| 30 | 213 | (** nat_rec -- used to define eclose and transrec, then obsolete; | 
| 214 | rec, from arith.ML, has fewer typing conditions **) | |
| 0 | 215 | |
| 216 | val nat_rec_trans = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); | |
| 217 | ||
| 5067 | 218 | Goal "nat_rec(0,a,b) = a"; | 
| 0 | 219 | by (rtac nat_rec_trans 1); | 
| 220 | by (rtac nat_case_0 1); | |
| 760 | 221 | qed "nat_rec_0"; | 
| 0 | 222 | |
| 223 | val [prem] = goal Nat.thy | |
| 224 | "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; | |
| 225 | by (rtac nat_rec_trans 1); | |
| 4091 | 226 | by (simp_tac (simpset() addsimps [prem, Memrel_iff, vimage_singleton_iff]) 1); | 
| 760 | 227 | qed "nat_rec_succ"; | 
| 0 | 228 | |
| 229 | (** The union of two natural numbers is a natural number -- their maximum **) | |
| 230 | ||
| 5137 | 231 | Goal "[| i: nat; j: nat |] ==> i Un j: nat"; | 
| 30 | 232 | by (rtac (Un_least_lt RS ltD) 1); | 
| 233 | by (REPEAT (ares_tac [ltI, Ord_nat] 1)); | |
| 760 | 234 | qed "Un_nat_type"; | 
| 0 | 235 | |
| 5137 | 236 | Goal "[| i: nat; j: nat |] ==> i Int j: nat"; | 
| 30 | 237 | by (rtac (Int_greatest_lt RS ltD) 1); | 
| 238 | by (REPEAT (ares_tac [ltI, Ord_nat] 1)); | |
| 760 | 239 | qed "Int_nat_type"; |