author | wenzelm |
Tue, 16 Oct 2001 00:32:01 +0200 | |
changeset 11790 | 42393a11642d |
parent 10216 | e928bdf62014 |
child 12552 | d2d2ab3f1f37 |
permissions | -rw-r--r-- |
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9180
diff
changeset
|
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 |
||
5067 | 9 |
Goal "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"; |
0 | 10 |
by (rtac bnd_monoI 1); |
11 |
by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2)); |
|
12 |
by (cut_facts_tac [infinity] 1); |
|
4243 | 13 |
by (Blast_tac 1); |
760 | 14 |
qed "nat_bnd_mono"; |
0 | 15 |
|
16 |
(* nat = {0} Un {succ(x). x:nat} *) |
|
10216 | 17 |
bind_thm ("nat_unfold", nat_bnd_mono RS (nat_def RS def_lfp_unfold)); |
0 | 18 |
|
19 |
(** Type checking of 0 and successor **) |
|
20 |
||
5067 | 21 |
Goal "0 : nat"; |
2033 | 22 |
by (stac nat_unfold 1); |
0 | 23 |
by (rtac (singletonI RS UnI1) 1); |
760 | 24 |
qed "nat_0I"; |
0 | 25 |
|
5137 | 26 |
Goal "n : nat ==> succ(n) : nat"; |
2033 | 27 |
by (stac nat_unfold 1); |
5137 | 28 |
by (etac (RepFunI RS UnI2) 1); |
760 | 29 |
qed "nat_succI"; |
0 | 30 |
|
5067 | 31 |
Goal "1 : nat"; |
0 | 32 |
by (rtac (nat_0I RS nat_succI) 1); |
760 | 33 |
qed "nat_1I"; |
0 | 34 |
|
5067 | 35 |
Goal "2 : nat"; |
829 | 36 |
by (rtac (nat_1I RS nat_succI) 1); |
37 |
qed "nat_2I"; |
|
38 |
||
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9180
diff
changeset
|
39 |
AddIffs [nat_0I, nat_1I, nat_2I]; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9180
diff
changeset
|
40 |
AddSIs [nat_succI]; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9180
diff
changeset
|
41 |
AddTCs [nat_0I, nat_1I, nat_2I, nat_succI]; |
2469 | 42 |
|
5067 | 43 |
Goal "bool <= nat"; |
6153 | 44 |
by (blast_tac (claset() addSEs [boolE]) 1); |
760 | 45 |
qed "bool_subset_nat"; |
0 | 46 |
|
9907 | 47 |
bind_thm ("bool_into_nat", bool_subset_nat RS subsetD); |
0 | 48 |
|
49 |
||
50 |
(** Injectivity properties and induction **) |
|
51 |
||
52 |
(*Mathematical induction*) |
|
9180 | 53 |
val major::prems = Goal |
0 | 54 |
"[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"; |
55 |
by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1); |
|
4243 | 56 |
by (blast_tac (claset() addIs prems) 1); |
760 | 57 |
qed "nat_induct"; |
0 | 58 |
|
59 |
(*Perform induction on n, then prove the n:nat subgoal using prems. *) |
|
60 |
fun nat_ind_tac a prems i = |
|
61 |
EVERY [res_inst_tac [("n",a)] nat_induct i, |
|
1461 | 62 |
rename_last_tac a ["1"] (i+2), |
63 |
ares_tac prems i]; |
|
0 | 64 |
|
9180 | 65 |
val major::prems = Goal |
0 | 66 |
"[| 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:
6
diff
changeset
|
67 |
by (rtac (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1); |
0 | 68 |
by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1 |
69 |
ORELSE ares_tac prems 1)); |
|
760 | 70 |
qed "natE"; |
0 | 71 |
|
9180 | 72 |
Goal "n: nat ==> Ord(n)"; |
73 |
by (nat_ind_tac "n" [] 1); |
|
0 | 74 |
by (REPEAT (ares_tac [Ord_0, Ord_succ] 1)); |
760 | 75 |
qed "nat_into_Ord"; |
0 | 76 |
|
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6153
diff
changeset
|
77 |
Addsimps [nat_into_Ord]; |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6153
diff
changeset
|
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:
6
diff
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 *) |
9907 | 100 |
bind_thm ("succ_natD", [succI1, asm_rl, Ord_nat] MRS Ord_trans); |
5137 | 101 |
AddSDs [succ_natD]; |
102 |
||
103 |
Goal "succ(n): nat <-> n: nat"; |
|
104 |
by (Blast_tac 1); |
|
105 |
qed "nat_succ_iff"; |
|
106 |
||
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9180
diff
changeset
|
107 |
AddIffs [Ord_nat, Limit_nat, nat_succ_iff]; |
2469 | 108 |
|
5137 | 109 |
Goal "Limit(i) ==> nat le i"; |
1461 | 110 |
by (rtac subset_imp_le 1); |
484 | 111 |
by (rtac subsetI 1); |
1461 | 112 |
by (etac nat_induct 1); |
4091 | 113 |
by (blast_tac (claset() addIs [Limit_has_succ RS ltD, ltI, Limit_is_Ord]) 2); |
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9180
diff
changeset
|
114 |
by (REPEAT (ares_tac [Limit_has_0 RS ltD, Ord_nat, Limit_is_Ord] 1)); |
760 | 115 |
qed "nat_le_Limit"; |
484 | 116 |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
117 |
(* [| succ(i): k; k: nat |] ==> i: k *) |
9907 | 118 |
bind_thm ("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:
6
diff
changeset
|
119 |
|
5137 | 120 |
Goal "[| m<n; n: nat |] ==> m: nat"; |
30 | 121 |
by (etac ltE 1); |
122 |
by (etac (Ord_nat RSN (3,Ord_trans)) 1); |
|
123 |
by (assume_tac 1); |
|
760 | 124 |
qed "lt_nat_in_nat"; |
30 | 125 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5147
diff
changeset
|
126 |
Goal "[| m le n; n:nat |] ==> m:nat"; |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5147
diff
changeset
|
127 |
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:
5147
diff
changeset
|
128 |
qed "le_in_nat"; |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5147
diff
changeset
|
129 |
|
30 | 130 |
|
0 | 131 |
(** Variations on mathematical induction **) |
132 |
||
133 |
(*complete induction*) |
|
9907 | 134 |
bind_thm ("complete_induct", Ord_nat RSN (2, Ord_induct)); |
0 | 135 |
|
9180 | 136 |
val prems = Goal |
0 | 137 |
"[| m: nat; n: nat; \ |
30 | 138 |
\ !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) \ |
139 |
\ |] ==> m le n --> P(m) --> P(n)"; |
|
0 | 140 |
by (nat_ind_tac "n" prems 1); |
141 |
by (ALLGOALS |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
142 |
(asm_simp_tac |
5529 | 143 |
(simpset() addsimps prems @ distrib_simps @ [le0_iff, le_succ_iff]))); |
760 | 144 |
qed "nat_induct_from_lemma"; |
0 | 145 |
|
146 |
(*Induction starting from m rather than 0*) |
|
9180 | 147 |
val prems = Goal |
30 | 148 |
"[| m le n; m: nat; n: nat; \ |
0 | 149 |
\ P(m); \ |
30 | 150 |
\ !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) \ |
0 | 151 |
\ |] ==> P(n)"; |
152 |
by (rtac (nat_induct_from_lemma RS mp RS mp) 1); |
|
153 |
by (REPEAT (ares_tac prems 1)); |
|
760 | 154 |
qed "nat_induct_from"; |
0 | 155 |
|
156 |
(*Induction suitable for subtraction and less-than*) |
|
9180 | 157 |
val prems = Goal |
0 | 158 |
"[| m: nat; n: nat; \ |
30 | 159 |
\ !!x. x: nat ==> P(x,0); \ |
160 |
\ !!y. y: nat ==> P(0,succ(y)); \ |
|
0 | 161 |
\ !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) \ |
162 |
\ |] ==> P(m,n)"; |
|
163 |
by (res_inst_tac [("x","m")] bspec 1); |
|
164 |
by (resolve_tac prems 2); |
|
165 |
by (nat_ind_tac "n" prems 1); |
|
166 |
by (rtac ballI 2); |
|
167 |
by (nat_ind_tac "x" [] 2); |
|
168 |
by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1)); |
|
760 | 169 |
qed "diff_induct"; |
0 | 170 |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
171 |
(** Induction principle analogous to trancl_induct **) |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
172 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
173 |
Goal "m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> \ |
30 | 174 |
\ (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:
6
diff
changeset
|
175 |
by (etac nat_induct 1); |
5479 | 176 |
by (ALLGOALS |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
177 |
(EVERY' [rtac (impI RS impI), rtac (nat_induct RS ballI), assume_tac, |
5479 | 178 |
blast_tac le_cs, blast_tac le_cs])); |
760 | 179 |
qed "succ_lt_induct_lemma"; |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
180 |
|
9180 | 181 |
val prems = Goal |
1461 | 182 |
"[| m<n; n: nat; \ |
183 |
\ P(m,succ(m)); \ |
|
184 |
\ !!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:
6
diff
changeset
|
185 |
\ |] ==> P(m,n)"; |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
186 |
by (res_inst_tac [("P4","P")] |
30 | 187 |
(succ_lt_induct_lemma RS mp RS mp RS bspec RS mp) 1); |
188 |
by (REPEAT (ares_tac (prems @ [ballI, impI, lt_nat_in_nat]) 1)); |
|
760 | 189 |
qed "succ_lt_induct"; |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
190 |
|
0 | 191 |
(** nat_case **) |
192 |
||
5067 | 193 |
Goalw [nat_case_def] "nat_case(a,b,0) = a"; |
5505 | 194 |
by (Blast_tac 1); |
760 | 195 |
qed "nat_case_0"; |
0 | 196 |
|
5067 | 197 |
Goalw [nat_case_def] "nat_case(a,b,succ(m)) = b(m)"; |
5505 | 198 |
by (Blast_tac 1); |
760 | 199 |
qed "nat_case_succ"; |
0 | 200 |
|
2469 | 201 |
Addsimps [nat_case_0, nat_case_succ]; |
202 |
||
9180 | 203 |
val major::prems = Goal |
0 | 204 |
"[| 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:
0
diff
changeset
|
205 |
\ |] ==> nat_case(a,b,n) : C(n)"; |
0 | 206 |
by (rtac (major RS nat_induct) 1); |
4091 | 207 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); |
760 | 208 |
qed "nat_case_type"; |
0 | 209 |
|
210 |
||
30 | 211 |
(** nat_rec -- used to define eclose and transrec, then obsolete; |
212 |
rec, from arith.ML, has fewer typing conditions **) |
|
0 | 213 |
|
6070 | 214 |
val lemma = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); |
0 | 215 |
|
5067 | 216 |
Goal "nat_rec(0,a,b) = a"; |
6070 | 217 |
by (rtac lemma 1); |
0 | 218 |
by (rtac nat_case_0 1); |
760 | 219 |
qed "nat_rec_0"; |
0 | 220 |
|
6070 | 221 |
Goal "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; |
222 |
by (rtac lemma 1); |
|
9842 | 223 |
by (asm_simp_tac (simpset() addsimps [vimage_singleton_iff]) 1); |
760 | 224 |
qed "nat_rec_succ"; |
0 | 225 |
|
226 |
(** The union of two natural numbers is a natural number -- their maximum **) |
|
227 |
||
5137 | 228 |
Goal "[| i: nat; j: nat |] ==> i Un j: nat"; |
30 | 229 |
by (rtac (Un_least_lt RS ltD) 1); |
230 |
by (REPEAT (ares_tac [ltI, Ord_nat] 1)); |
|
760 | 231 |
qed "Un_nat_type"; |
0 | 232 |
|
5137 | 233 |
Goal "[| i: nat; j: nat |] ==> i Int j: nat"; |
30 | 234 |
by (rtac (Int_greatest_lt RS ltD) 1); |
235 |
by (REPEAT (ares_tac [ltI, Ord_nat] 1)); |
|
760 | 236 |
qed "Int_nat_type"; |