author | wenzelm |
Sun, 25 Oct 1998 12:33:27 +0100 | |
changeset 5769 | 6a422b22ba02 |
parent 5268 | 59ef39008514 |
child 6044 | e0f9d930e956 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/ex/Primrec |
515 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
515 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Primitive Recursive Functions |
|
7 |
||
8 |
Proof adopted from |
|
9 |
Nora Szasz, |
|
10 |
A Machine Checked Proof that Ackermann's Function is not Primitive Recursive, |
|
11 |
In: Huet & Plotkin, eds., Logical Environments (CUP, 1993), 317-338. |
|
12 |
||
13 |
See also E. Mendelson, Introduction to Mathematical Logic. |
|
14 |
(Van Nostrand, 1964), page 250, exercise 11. |
|
15 |
*) |
|
16 |
||
17 |
open Primrec; |
|
18 |
||
19 |
val pr_typechecks = |
|
20 |
nat_typechecks @ list.intrs @ |
|
21 |
[lam_type, list_case_type, drop_type, map_type, apply_type, rec_type]; |
|
22 |
||
23 |
(** Useful special cases of evaluation ***) |
|
24 |
||
4091 | 25 |
simpset_ref() := simpset() setSolver (type_auto_tac pr_typechecks); |
515 | 26 |
|
5068 | 27 |
Goalw [SC_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
28 |
"[| x:nat; l: list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)"; |
2469 | 29 |
by (Asm_simp_tac 1); |
760 | 30 |
qed "SC"; |
515 | 31 |
|
5068 | 32 |
Goalw [CONST_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
33 |
"[| l: list(nat) |] ==> CONST(k) ` l = k"; |
2469 | 34 |
by (Asm_simp_tac 1); |
760 | 35 |
qed "CONST"; |
515 | 36 |
|
5068 | 37 |
Goalw [PROJ_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
38 |
"[| x: nat; l: list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x"; |
2469 | 39 |
by (Asm_simp_tac 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
40 |
qed "PROJ_0"; |
515 | 41 |
|
5068 | 42 |
Goalw [COMP_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
43 |
"[| l: list(nat) |] ==> COMP(g,[f]) ` l = g` [f`l]"; |
2469 | 44 |
by (Asm_simp_tac 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
45 |
qed "COMP_1"; |
515 | 46 |
|
5068 | 47 |
Goalw [PREC_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
48 |
"l: list(nat) ==> PREC(f,g) ` (Cons(0,l)) = f`l"; |
2469 | 49 |
by (Asm_simp_tac 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
50 |
qed "PREC_0"; |
515 | 51 |
|
5068 | 52 |
Goalw [PREC_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
53 |
"[| x:nat; l: list(nat) |] ==> \ |
515 | 54 |
\ PREC(f,g) ` (Cons(succ(x),l)) = \ |
55 |
\ g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))"; |
|
2469 | 56 |
by (Asm_simp_tac 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
57 |
qed "PREC_succ"; |
515 | 58 |
|
59 |
(*** Inductive definition of the PR functions ***) |
|
60 |
||
61 |
(* c: primrec ==> c: list(nat) -> nat *) |
|
62 |
val primrec_into_fun = primrec.dom_subset RS subsetD; |
|
63 |
||
4091 | 64 |
simpset_ref() := simpset() setSolver (type_auto_tac ([primrec_into_fun] @ |
2469 | 65 |
pr_typechecks @ primrec.intrs)); |
515 | 66 |
|
5137 | 67 |
Goalw [ACK_def] "i:nat ==> ACK(i): primrec"; |
515 | 68 |
by (etac nat_induct 1); |
2469 | 69 |
by (ALLGOALS Asm_simp_tac); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
70 |
qed "ACK_in_primrec"; |
515 | 71 |
|
72 |
val ack_typechecks = |
|
73 |
[ACK_in_primrec, primrec_into_fun RS apply_type, |
|
74 |
add_type, list_add_type, nat_into_Ord] @ |
|
75 |
nat_typechecks @ list.intrs @ primrec.intrs; |
|
76 |
||
77 |
(*strict typechecking for the Ackermann proof; instantiates no vars*) |
|
78 |
fun tc_tac rls = |
|
79 |
REPEAT |
|
80 |
(SOMEGOAL (test_assume_tac ORELSE' match_tac (rls @ ack_typechecks))); |
|
81 |
||
5137 | 82 |
Goal "[| i:nat; j:nat |] ==> ack(i,j): nat"; |
515 | 83 |
by (tc_tac []); |
760 | 84 |
qed "ack_type"; |
515 | 85 |
|
86 |
(** Ackermann's function cases **) |
|
87 |
||
88 |
(*PROPERTY A 1*) |
|
5137 | 89 |
Goalw [ACK_def] "j:nat ==> ack(0,j) = succ(j)"; |
4091 | 90 |
by (asm_simp_tac (simpset() addsimps [SC]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
91 |
qed "ack_0"; |
515 | 92 |
|
93 |
(*PROPERTY A 2*) |
|
5068 | 94 |
Goalw [ACK_def] "ack(succ(i), 0) = ack(i,1)"; |
4091 | 95 |
by (asm_simp_tac (simpset() addsimps [CONST,PREC_0]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
96 |
qed "ack_succ_0"; |
515 | 97 |
|
98 |
(*PROPERTY A 3*) |
|
99 |
(*Could be proved in Primrec0, like the previous two cases, but using |
|
100 |
primrec_into_fun makes type-checking easier!*) |
|
5068 | 101 |
Goalw [ACK_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
102 |
"[| i:nat; j:nat |] ==> \ |
515 | 103 |
\ ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))"; |
4091 | 104 |
by (asm_simp_tac (simpset() addsimps [CONST,PREC_succ,COMP_1,PROJ_0]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
105 |
qed "ack_succ_succ"; |
515 | 106 |
|
2469 | 107 |
Addsimps [ack_0, ack_succ_0, ack_succ_succ, ack_type, nat_into_Ord]; |
515 | 108 |
|
109 |
(*PROPERTY A 4*) |
|
5137 | 110 |
Goal "i:nat ==> ALL j:nat. j < ack(i,j)"; |
515 | 111 |
by (etac nat_induct 1); |
2469 | 112 |
by (Asm_simp_tac 1); |
515 | 113 |
by (rtac ballI 1); |
114 |
by (eres_inst_tac [("n","j")] nat_induct 1); |
|
115 |
by (DO_GOAL [rtac (nat_0I RS nat_0_le RS lt_trans), |
|
2469 | 116 |
Asm_simp_tac] 1); |
515 | 117 |
by (DO_GOAL [etac (succ_leI RS lt_trans1), |
2469 | 118 |
Asm_simp_tac] 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
119 |
qed "lt_ack2_lemma"; |
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
120 |
bind_thm ("lt_ack2", (lt_ack2_lemma RS bspec)); |
515 | 121 |
|
122 |
(*PROPERTY A 5-, the single-step lemma*) |
|
5137 | 123 |
Goal "[| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))"; |
515 | 124 |
by (etac nat_induct 1); |
4091 | 125 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [lt_ack2]))); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
126 |
qed "ack_lt_ack_succ2"; |
515 | 127 |
|
128 |
(*PROPERTY A 5, monotonicity for < *) |
|
5137 | 129 |
Goal "[| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)"; |
515 | 130 |
by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1); |
131 |
by (etac succ_lt_induct 1); |
|
132 |
by (assume_tac 1); |
|
133 |
by (rtac lt_trans 2); |
|
134 |
by (REPEAT (ares_tac ([ack_lt_ack_succ2, ack_type] @ pr_typechecks) 1)); |
|
760 | 135 |
qed "ack_lt_mono2"; |
515 | 136 |
|
137 |
(*PROPERTY A 5', monotonicity for le *) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
138 |
Goal "[| j le k; i: nat; k:nat |] ==> ack(i,j) le ack(i,k)"; |
3840 | 139 |
by (res_inst_tac [("f", "%j. ack(i,j)")] Ord_lt_mono_imp_le_mono 1); |
515 | 140 |
by (REPEAT (ares_tac [ack_lt_mono2, ack_type RS nat_into_Ord] 1)); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
141 |
qed "ack_le_mono2"; |
515 | 142 |
|
143 |
(*PROPERTY A 6*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
144 |
Goal "[| i:nat; j:nat |] ==> ack(i, succ(j)) le ack(succ(i), j)"; |
515 | 145 |
by (nat_ind_tac "j" [] 1); |
2469 | 146 |
by (ALLGOALS Asm_simp_tac); |
515 | 147 |
by (rtac ack_le_mono2 1); |
148 |
by (rtac (lt_ack2 RS succ_leI RS le_trans) 1); |
|
149 |
by (REPEAT (ares_tac (ack_typechecks) 1)); |
|
760 | 150 |
qed "ack2_le_ack1"; |
515 | 151 |
|
152 |
(*PROPERTY A 7-, the single-step lemma*) |
|
5137 | 153 |
Goal "[| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)"; |
515 | 154 |
by (rtac (ack_lt_mono2 RS lt_trans2) 1); |
155 |
by (rtac ack2_le_ack1 4); |
|
156 |
by (REPEAT (ares_tac ([nat_le_refl, ack_type] @ pr_typechecks) 1)); |
|
760 | 157 |
qed "ack_lt_ack_succ1"; |
515 | 158 |
|
159 |
(*PROPERTY A 7, monotonicity for < *) |
|
5137 | 160 |
Goal "[| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)"; |
515 | 161 |
by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1); |
162 |
by (etac succ_lt_induct 1); |
|
163 |
by (assume_tac 1); |
|
164 |
by (rtac lt_trans 2); |
|
165 |
by (REPEAT (ares_tac ([ack_lt_ack_succ1, ack_type] @ pr_typechecks) 1)); |
|
760 | 166 |
qed "ack_lt_mono1"; |
515 | 167 |
|
168 |
(*PROPERTY A 7', monotonicity for le *) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
169 |
Goal "[| i le j; j:nat; k:nat |] ==> ack(i,k) le ack(j,k)"; |
3840 | 170 |
by (res_inst_tac [("f", "%j. ack(j,k)")] Ord_lt_mono_imp_le_mono 1); |
515 | 171 |
by (REPEAT (ares_tac [ack_lt_mono1, ack_type RS nat_into_Ord] 1)); |
760 | 172 |
qed "ack_le_mono1"; |
515 | 173 |
|
174 |
(*PROPERTY A 8*) |
|
5137 | 175 |
Goal "j:nat ==> ack(1,j) = succ(succ(j))"; |
515 | 176 |
by (etac nat_induct 1); |
2469 | 177 |
by (ALLGOALS Asm_simp_tac); |
760 | 178 |
qed "ack_1"; |
515 | 179 |
|
180 |
(*PROPERTY A 9*) |
|
5137 | 181 |
Goal "j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))"; |
515 | 182 |
by (etac nat_induct 1); |
4091 | 183 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [ack_1, add_succ_right]))); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
184 |
qed "ack_2"; |
515 | 185 |
|
186 |
(*PROPERTY A 10*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
187 |
Goal "[| i1:nat; i2:nat; j:nat |] ==> \ |
515 | 188 |
\ ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)"; |
189 |
by (rtac (ack2_le_ack1 RSN (2,lt_trans2)) 1); |
|
2469 | 190 |
by (Asm_simp_tac 1); |
515 | 191 |
by (rtac (add_le_self RS ack_le_mono1 RS lt_trans1) 1); |
192 |
by (rtac (add_le_self2 RS ack_lt_mono1 RS ack_lt_mono2) 5); |
|
193 |
by (tc_tac []); |
|
760 | 194 |
qed "ack_nest_bound"; |
515 | 195 |
|
196 |
(*PROPERTY A 11*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
197 |
Goal "[| i1:nat; i2:nat; j:nat |] ==> \ |
515 | 198 |
\ ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)"; |
199 |
by (res_inst_tac [("j", "ack(succ(1), ack(i1 #+ i2, j))")] lt_trans 1); |
|
4091 | 200 |
by (asm_simp_tac (simpset() addsimps [ack_2]) 1); |
515 | 201 |
by (rtac (ack_nest_bound RS lt_trans2) 2); |
2469 | 202 |
by (Asm_simp_tac 5); |
515 | 203 |
by (rtac (add_le_mono RS leI RS leI) 1); |
204 |
by (REPEAT (ares_tac ([add_le_self, add_le_self2, ack_le_mono1] @ |
|
205 |
ack_typechecks) 1)); |
|
760 | 206 |
qed "ack_add_bound"; |
515 | 207 |
|
208 |
(*PROPERTY A 12. Article uses existential quantifier but the ALF proof |
|
209 |
used k#+4. Quantified version must be nested EX k'. ALL i,j... *) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
210 |
Goal "[| i < ack(k,j); j:nat; k:nat |] ==> \ |
515 | 211 |
\ i#+j < ack(succ(succ(succ(succ(k)))), j)"; |
212 |
by (res_inst_tac [("j", "ack(k,j) #+ ack(0,j)")] lt_trans 1); |
|
213 |
by (rtac (ack_add_bound RS lt_trans2) 2); |
|
4091 | 214 |
by (asm_simp_tac (simpset() addsimps [add_0_right]) 5); |
515 | 215 |
by (REPEAT (ares_tac ([add_lt_mono, lt_ack2] @ ack_typechecks) 1)); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
216 |
qed "ack_add_bound2"; |
515 | 217 |
|
218 |
(*** MAIN RESULT ***) |
|
219 |
||
2469 | 220 |
Addsimps [list_add_type, nat_into_Ord]; |
515 | 221 |
|
5068 | 222 |
Goalw [SC_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
223 |
"l: list(nat) ==> SC ` l < ack(1, list_add(l))"; |
515 | 224 |
by (etac list.elim 1); |
4091 | 225 |
by (asm_simp_tac (simpset() addsimps [succ_iff]) 1); |
226 |
by (asm_simp_tac (simpset() addsimps [ack_1, add_le_self]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
227 |
qed "SC_case"; |
515 | 228 |
|
229 |
(*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*) |
|
5137 | 230 |
Goal "[| i:nat; j:nat |] ==> i < ack(i,j)"; |
515 | 231 |
by (etac nat_induct 1); |
4091 | 232 |
by (asm_simp_tac (simpset() addsimps [nat_0_le]) 1); |
515 | 233 |
by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1); |
234 |
by (tc_tac []); |
|
760 | 235 |
qed "lt_ack1"; |
515 | 236 |
|
5068 | 237 |
Goalw [CONST_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
238 |
"[| l: list(nat); k: nat |] ==> CONST(k) ` l < ack(k, list_add(l))"; |
4091 | 239 |
by (asm_simp_tac (simpset() addsimps [lt_ack1]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
240 |
qed "CONST_case"; |
515 | 241 |
|
5068 | 242 |
Goalw [PROJ_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
243 |
"l: list(nat) ==> ALL i:nat. PROJ(i) ` l < ack(0, list_add(l))"; |
2469 | 244 |
by (Asm_simp_tac 1); |
515 | 245 |
by (etac list.induct 1); |
4091 | 246 |
by (asm_simp_tac (simpset() addsimps [nat_0_le]) 1); |
2469 | 247 |
by (Asm_simp_tac 1); |
515 | 248 |
by (rtac ballI 1); |
249 |
by (eres_inst_tac [("n","x")] natE 1); |
|
4091 | 250 |
by (asm_simp_tac (simpset() addsimps [add_le_self]) 1); |
2469 | 251 |
by (Asm_simp_tac 1); |
515 | 252 |
by (etac (bspec RS lt_trans2) 1); |
253 |
by (rtac (add_le_self2 RS succ_leI) 2); |
|
254 |
by (tc_tac []); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
255 |
qed "PROJ_case_lemma"; |
515 | 256 |
val PROJ_case = PROJ_case_lemma RS bspec; |
257 |
||
258 |
(** COMP case **) |
|
259 |
||
5268 | 260 |
Goal "fs : list({f: primrec . \ |
1461 | 261 |
\ EX kf:nat. ALL l:list(nat). \ |
262 |
\ f`l < ack(kf, list_add(l))}) \ |
|
263 |
\ ==> EX k:nat. ALL l: list(nat). \ |
|
515 | 264 |
\ list_add(map(%f. f ` l, fs)) < ack(k, list_add(l))"; |
265 |
by (etac list.induct 1); |
|
266 |
by (DO_GOAL [res_inst_tac [("x","0")] bexI, |
|
4091 | 267 |
asm_simp_tac (simpset() addsimps [lt_ack1, nat_0_le]), |
1461 | 268 |
resolve_tac nat_typechecks] 1); |
4152 | 269 |
by Safe_tac; |
2469 | 270 |
by (Asm_simp_tac 1); |
515 | 271 |
by (rtac (ballI RS bexI) 1); |
272 |
by (rtac (add_lt_mono RS lt_trans) 1); |
|
273 |
by (REPEAT (FIRSTGOAL (etac bspec))); |
|
274 |
by (rtac ack_add_bound 5); |
|
275 |
by (tc_tac []); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
276 |
qed "COMP_map_lemma"; |
515 | 277 |
|
5068 | 278 |
Goalw [COMP_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
279 |
"[| kg: nat; \ |
1461 | 280 |
\ ALL l:list(nat). g`l < ack(kg, list_add(l)); \ |
281 |
\ fs : list({f: primrec . \ |
|
282 |
\ EX kf:nat. ALL l:list(nat). \ |
|
283 |
\ f`l < ack(kf, list_add(l))}) \ |
|
515 | 284 |
\ |] ==> EX k:nat. ALL l: list(nat). COMP(g,fs)`l < ack(k, list_add(l))"; |
2469 | 285 |
by (Asm_simp_tac 1); |
515 | 286 |
by (forward_tac [list_CollectD] 1); |
287 |
by (etac (COMP_map_lemma RS bexE) 1); |
|
288 |
by (rtac (ballI RS bexI) 1); |
|
289 |
by (etac (bspec RS lt_trans) 1); |
|
290 |
by (rtac lt_trans 2); |
|
291 |
by (rtac ack_nest_bound 3); |
|
292 |
by (etac (bspec RS ack_lt_mono2) 2); |
|
293 |
by (tc_tac [map_type]); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
294 |
qed "COMP_case"; |
515 | 295 |
|
296 |
(** PREC case **) |
|
297 |
||
5068 | 298 |
Goalw [PREC_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
299 |
"[| ALL l:list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \ |
1461 | 300 |
\ ALL l:list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); \ |
301 |
\ f: primrec; kf: nat; \ |
|
302 |
\ g: primrec; kg: nat; \ |
|
303 |
\ l: list(nat) \ |
|
515 | 304 |
\ |] ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))"; |
305 |
by (etac list.elim 1); |
|
4091 | 306 |
by (asm_simp_tac (simpset() addsimps [[nat_le_refl, lt_ack2] MRS lt_trans]) 1); |
2469 | 307 |
by (Asm_simp_tac 1); |
515 | 308 |
by (etac ssubst 1); (*get rid of the needless assumption*) |
309 |
by (eres_inst_tac [("n","a")] nat_induct 1); |
|
310 |
(*base case*) |
|
2469 | 311 |
by (DO_GOAL [Asm_simp_tac, rtac lt_trans, etac bspec, |
1461 | 312 |
assume_tac, rtac (add_le_self RS ack_lt_mono1), |
313 |
REPEAT o ares_tac (ack_typechecks)] 1); |
|
515 | 314 |
(*ind step*) |
3328 | 315 |
by (Asm_simp_tac 1); |
515 | 316 |
by (rtac (succ_leI RS lt_trans1) 1); |
317 |
by (res_inst_tac [("j", "g ` ?ll #+ ?mm")] lt_trans1 1); |
|
318 |
by (etac bspec 2); |
|
319 |
by (rtac (nat_le_refl RS add_le_mono) 1); |
|
320 |
by (tc_tac []); |
|
4091 | 321 |
by (asm_simp_tac (simpset() addsimps [add_le_self2]) 1); |
515 | 322 |
(*final part of the simplification*) |
2469 | 323 |
by (Asm_simp_tac 1); |
515 | 324 |
by (rtac (add_le_self2 RS ack_le_mono1 RS lt_trans1) 1); |
325 |
by (etac ack_lt_mono2 5); |
|
326 |
by (tc_tac []); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
327 |
qed "PREC_case_lemma"; |
515 | 328 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
329 |
Goal "[| f: primrec; kf: nat; \ |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
330 |
\ g: primrec; kg: nat; \ |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
331 |
\ ALL l:list(nat). f`l < ack(kf, list_add(l)); \ |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
332 |
\ ALL l:list(nat). g`l < ack(kg, list_add(l)) \ |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
333 |
\ |] ==> EX k:nat. ALL l: list(nat). PREC(f,g)`l< ack(k, list_add(l))"; |
515 | 334 |
by (rtac (ballI RS bexI) 1); |
335 |
by (rtac ([add_le_self, PREC_case_lemma] MRS lt_trans1) 1); |
|
336 |
by (REPEAT |
|
337 |
(SOMEGOAL |
|
338 |
(FIRST' [test_assume_tac, |
|
1461 | 339 |
match_tac (ack_typechecks), |
340 |
rtac (ack_add_bound2 RS ballI) THEN' etac bspec]))); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
341 |
qed "PREC_case"; |
515 | 342 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
343 |
Goal "f:primrec ==> EX k:nat. ALL l:list(nat). f`l < ack(k, list_add(l))"; |
515 | 344 |
by (etac primrec.induct 1); |
4152 | 345 |
by Safe_tac; |
515 | 346 |
by (DEPTH_SOLVE |
347 |
(ares_tac ([SC_case, CONST_case, PROJ_case, COMP_case, PREC_case, |
|
1461 | 348 |
bexI, ballI] @ nat_typechecks) 1)); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
349 |
qed "ack_bounds_primrec"; |
515 | 350 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
351 |
Goal "~ (lam l:list(nat). list_case(0, %x xs. ack(x,x), l)) : primrec"; |
515 | 352 |
by (rtac notI 1); |
353 |
by (etac (ack_bounds_primrec RS bexE) 1); |
|
354 |
by (rtac lt_irrefl 1); |
|
355 |
by (dres_inst_tac [("x", "[x]")] bspec 1); |
|
2469 | 356 |
by (Asm_simp_tac 1); |
4091 | 357 |
by (asm_full_simp_tac (simpset() addsimps [add_0_right]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
358 |
qed "ack_not_primrec"; |
515 | 359 |