| author | wenzelm | 
| Fri, 31 Oct 1997 15:21:32 +0100 | |
| changeset 4054 | b33e02b3478e | 
| parent 3840 | e0baea4d485a | 
| child 4091 | 771b1f6422a8 | 
| 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  | 
||
| 
2637
 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 
oheimb 
parents: 
2469 
diff
changeset
 | 
25  | 
simpset := !simpset setSolver (type_auto_tac pr_typechecks);  | 
| 515 | 26  | 
|
27  | 
goalw Primrec.thy [SC_def]  | 
|
28  | 
"!!x l. [| x:nat; l: list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)";  | 
|
| 2469 | 29  | 
by (Asm_simp_tac 1);  | 
| 760 | 30  | 
qed "SC";  | 
| 515 | 31  | 
|
32  | 
goalw Primrec.thy [CONST_def]  | 
|
33  | 
"!!l. [| l: list(nat) |] ==> CONST(k) ` l = k";  | 
|
| 2469 | 34  | 
by (Asm_simp_tac 1);  | 
| 760 | 35  | 
qed "CONST";  | 
| 515 | 36  | 
|
37  | 
goalw Primrec.thy [PROJ_def]  | 
|
38  | 
"!!l. [| 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  | 
|
42  | 
goalw Primrec.thy [COMP_def]  | 
|
43  | 
"!!l. [| 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  | 
|
47  | 
goalw Primrec.thy [PREC_def]  | 
|
48  | 
"!!l. 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  | 
|
52  | 
goalw Primrec.thy [PREC_def]  | 
|
53  | 
"!!l. [| x:nat; l: list(nat) |] ==> \  | 
|
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  | 
||
| 
2637
 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 
oheimb 
parents: 
2469 
diff
changeset
 | 
64  | 
simpset := !simpset setSolver (type_auto_tac ([primrec_into_fun] @  | 
| 2469 | 65  | 
pr_typechecks @ primrec.intrs));  | 
| 515 | 66  | 
|
67  | 
goalw Primrec.thy [ACK_def] "!!i. i:nat ==> ACK(i): primrec";  | 
|
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  | 
||
82  | 
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j): nat";  | 
|
83  | 
by (tc_tac []);  | 
|
| 760 | 84  | 
qed "ack_type";  | 
| 515 | 85  | 
|
86  | 
(** Ackermann's function cases **)  | 
|
87  | 
||
88  | 
(*PROPERTY A 1*)  | 
|
89  | 
goalw Primrec.thy [ACK_def] "!!j. j:nat ==> ack(0,j) = succ(j)";  | 
|
| 2469 | 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*)  | 
|
94  | 
goalw Primrec.thy [ACK_def] "ack(succ(i), 0) = ack(i,1)";  | 
|
| 2469 | 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!*)  | 
|
101  | 
goalw Primrec.thy [ACK_def]  | 
|
102  | 
"!!i j. [| i:nat; j:nat |] ==> \  | 
|
103  | 
\ ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))";  | 
|
| 2469 | 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*)  | 
|
110  | 
goal Primrec.thy "!!i. i:nat ==> ALL j:nat. j < ack(i,j)";  | 
|
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*)  | 
|
123  | 
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";  | 
|
124  | 
by (etac nat_induct 1);  | 
|
| 2469 | 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 < *)  | 
|
129  | 
goal Primrec.thy "!!i j k. [| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)";  | 
|
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 *)  | 
|
138  | 
goal Primrec.thy  | 
|
139  | 
"!!i j k. [| j le k; i: nat; k:nat |] ==> ack(i,j) le ack(i,k)";  | 
|
| 3840 | 140  | 
by (res_inst_tac [("f", "%j. ack(i,j)")] Ord_lt_mono_imp_le_mono 1);
 | 
| 515 | 141  | 
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
 | 
142  | 
qed "ack_le_mono2";  | 
| 515 | 143  | 
|
144  | 
(*PROPERTY A 6*)  | 
|
145  | 
goal Primrec.thy  | 
|
146  | 
"!!i j. [| i:nat; j:nat |] ==> ack(i, succ(j)) le ack(succ(i), j)";  | 
|
147  | 
by (nat_ind_tac "j" [] 1);  | 
|
| 2469 | 148  | 
by (ALLGOALS Asm_simp_tac);  | 
| 515 | 149  | 
by (rtac ack_le_mono2 1);  | 
150  | 
by (rtac (lt_ack2 RS succ_leI RS le_trans) 1);  | 
|
151  | 
by (REPEAT (ares_tac (ack_typechecks) 1));  | 
|
| 760 | 152  | 
qed "ack2_le_ack1";  | 
| 515 | 153  | 
|
154  | 
(*PROPERTY A 7-, the single-step lemma*)  | 
|
155  | 
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";  | 
|
156  | 
by (rtac (ack_lt_mono2 RS lt_trans2) 1);  | 
|
157  | 
by (rtac ack2_le_ack1 4);  | 
|
158  | 
by (REPEAT (ares_tac ([nat_le_refl, ack_type] @ pr_typechecks) 1));  | 
|
| 760 | 159  | 
qed "ack_lt_ack_succ1";  | 
| 515 | 160  | 
|
161  | 
(*PROPERTY A 7, monotonicity for < *)  | 
|
162  | 
goal Primrec.thy "!!i j k. [| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";  | 
|
163  | 
by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);  | 
|
164  | 
by (etac succ_lt_induct 1);  | 
|
165  | 
by (assume_tac 1);  | 
|
166  | 
by (rtac lt_trans 2);  | 
|
167  | 
by (REPEAT (ares_tac ([ack_lt_ack_succ1, ack_type] @ pr_typechecks) 1));  | 
|
| 760 | 168  | 
qed "ack_lt_mono1";  | 
| 515 | 169  | 
|
170  | 
(*PROPERTY A 7', monotonicity for le *)  | 
|
171  | 
goal Primrec.thy  | 
|
172  | 
"!!i j k. [| i le j; j:nat; k:nat |] ==> ack(i,k) le ack(j,k)";  | 
|
| 3840 | 173  | 
by (res_inst_tac [("f", "%j. ack(j,k)")] Ord_lt_mono_imp_le_mono 1);
 | 
| 515 | 174  | 
by (REPEAT (ares_tac [ack_lt_mono1, ack_type RS nat_into_Ord] 1));  | 
| 760 | 175  | 
qed "ack_le_mono1";  | 
| 515 | 176  | 
|
177  | 
(*PROPERTY A 8*)  | 
|
178  | 
goal Primrec.thy "!!j. j:nat ==> ack(1,j) = succ(succ(j))";  | 
|
179  | 
by (etac nat_induct 1);  | 
|
| 2469 | 180  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 181  | 
qed "ack_1";  | 
| 515 | 182  | 
|
183  | 
(*PROPERTY A 9*)  | 
|
184  | 
goal Primrec.thy "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";  | 
|
185  | 
by (etac nat_induct 1);  | 
|
| 2469 | 186  | 
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
 | 
187  | 
qed "ack_2";  | 
| 515 | 188  | 
|
189  | 
(*PROPERTY A 10*)  | 
|
190  | 
goal Primrec.thy  | 
|
191  | 
"!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \  | 
|
192  | 
\ ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)";  | 
|
193  | 
by (rtac (ack2_le_ack1 RSN (2,lt_trans2)) 1);  | 
|
| 2469 | 194  | 
by (Asm_simp_tac 1);  | 
| 515 | 195  | 
by (rtac (add_le_self RS ack_le_mono1 RS lt_trans1) 1);  | 
196  | 
by (rtac (add_le_self2 RS ack_lt_mono1 RS ack_lt_mono2) 5);  | 
|
197  | 
by (tc_tac []);  | 
|
| 760 | 198  | 
qed "ack_nest_bound";  | 
| 515 | 199  | 
|
200  | 
(*PROPERTY A 11*)  | 
|
201  | 
goal Primrec.thy  | 
|
202  | 
"!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \  | 
|
203  | 
\ ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)";  | 
|
204  | 
by (res_inst_tac [("j", "ack(succ(1), ack(i1 #+ i2, j))")] lt_trans 1);
 | 
|
| 2469 | 205  | 
by (asm_simp_tac (!simpset addsimps [ack_2]) 1);  | 
| 515 | 206  | 
by (rtac (ack_nest_bound RS lt_trans2) 2);  | 
| 2469 | 207  | 
by (Asm_simp_tac 5);  | 
| 515 | 208  | 
by (rtac (add_le_mono RS leI RS leI) 1);  | 
209  | 
by (REPEAT (ares_tac ([add_le_self, add_le_self2, ack_le_mono1] @  | 
|
210  | 
ack_typechecks) 1));  | 
|
| 760 | 211  | 
qed "ack_add_bound";  | 
| 515 | 212  | 
|
213  | 
(*PROPERTY A 12. Article uses existential quantifier but the ALF proof  | 
|
214  | 
used k#+4. Quantified version must be nested EX k'. ALL i,j... *)  | 
|
215  | 
goal Primrec.thy  | 
|
216  | 
"!!i j k. [| i < ack(k,j); j:nat; k:nat |] ==> \  | 
|
217  | 
\ i#+j < ack(succ(succ(succ(succ(k)))), j)";  | 
|
218  | 
by (res_inst_tac [("j", "ack(k,j) #+ ack(0,j)")] lt_trans 1);
 | 
|
219  | 
by (rtac (ack_add_bound RS lt_trans2) 2);  | 
|
| 2469 | 220  | 
by (asm_simp_tac (!simpset addsimps [add_0_right]) 5);  | 
| 515 | 221  | 
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
 | 
222  | 
qed "ack_add_bound2";  | 
| 515 | 223  | 
|
224  | 
(*** MAIN RESULT ***)  | 
|
225  | 
||
| 2469 | 226  | 
Addsimps [list_add_type, nat_into_Ord];  | 
| 515 | 227  | 
|
228  | 
goalw Primrec.thy [SC_def]  | 
|
229  | 
"!!l. l: list(nat) ==> SC ` l < ack(1, list_add(l))";  | 
|
230  | 
by (etac list.elim 1);  | 
|
| 2469 | 231  | 
by (asm_simp_tac (!simpset addsimps [succ_iff]) 1);  | 
232  | 
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
 | 
233  | 
qed "SC_case";  | 
| 515 | 234  | 
|
235  | 
(*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*)  | 
|
236  | 
goal Primrec.thy "!!j. [| i:nat; j:nat |] ==> i < ack(i,j)";  | 
|
237  | 
by (etac nat_induct 1);  | 
|
| 2469 | 238  | 
by (asm_simp_tac (!simpset addsimps [nat_0_le]) 1);  | 
| 515 | 239  | 
by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1);  | 
240  | 
by (tc_tac []);  | 
|
| 760 | 241  | 
qed "lt_ack1";  | 
| 515 | 242  | 
|
243  | 
goalw Primrec.thy [CONST_def]  | 
|
244  | 
"!!l. [| l: list(nat); k: nat |] ==> CONST(k) ` l < ack(k, list_add(l))";  | 
|
| 2469 | 245  | 
by (asm_simp_tac (!simpset addsimps [lt_ack1]) 1);  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
246  | 
qed "CONST_case";  | 
| 515 | 247  | 
|
248  | 
goalw Primrec.thy [PROJ_def]  | 
|
249  | 
"!!l. l: list(nat) ==> ALL i:nat. PROJ(i) ` l < ack(0, list_add(l))";  | 
|
| 2469 | 250  | 
by (Asm_simp_tac 1);  | 
| 515 | 251  | 
by (etac list.induct 1);  | 
| 2469 | 252  | 
by (asm_simp_tac (!simpset addsimps [nat_0_le]) 1);  | 
253  | 
by (Asm_simp_tac 1);  | 
|
| 515 | 254  | 
by (rtac ballI 1);  | 
255  | 
by (eres_inst_tac [("n","x")] natE 1);
 | 
|
| 2469 | 256  | 
by (asm_simp_tac (!simpset addsimps [add_le_self]) 1);  | 
257  | 
by (Asm_simp_tac 1);  | 
|
| 515 | 258  | 
by (etac (bspec RS lt_trans2) 1);  | 
259  | 
by (rtac (add_le_self2 RS succ_leI) 2);  | 
|
260  | 
by (tc_tac []);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
261  | 
qed "PROJ_case_lemma";  | 
| 515 | 262  | 
val PROJ_case = PROJ_case_lemma RS bspec;  | 
263  | 
||
264  | 
(** COMP case **)  | 
|
265  | 
||
266  | 
goal Primrec.thy  | 
|
| 1461 | 267  | 
 "!!fs. fs : list({f: primrec .                                 \
 | 
268  | 
\ EX kf:nat. ALL l:list(nat). \  | 
|
269  | 
\ f`l < ack(kf, list_add(l))}) \  | 
|
270  | 
\ ==> EX k:nat. ALL l: list(nat). \  | 
|
| 515 | 271  | 
\ list_add(map(%f. f ` l, fs)) < ack(k, list_add(l))";  | 
272  | 
by (etac list.induct 1);  | 
|
273  | 
by (DO_GOAL [res_inst_tac [("x","0")] bexI,
 | 
|
| 2469 | 274  | 
asm_simp_tac (!simpset addsimps [lt_ack1, nat_0_le]),  | 
| 1461 | 275  | 
resolve_tac nat_typechecks] 1);  | 
| 2469 | 276  | 
by (safe_tac (!claset));  | 
277  | 
by (Asm_simp_tac 1);  | 
|
| 515 | 278  | 
by (rtac (ballI RS bexI) 1);  | 
279  | 
by (rtac (add_lt_mono RS lt_trans) 1);  | 
|
280  | 
by (REPEAT (FIRSTGOAL (etac bspec)));  | 
|
281  | 
by (rtac ack_add_bound 5);  | 
|
282  | 
by (tc_tac []);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
283  | 
qed "COMP_map_lemma";  | 
| 515 | 284  | 
|
285  | 
goalw Primrec.thy [COMP_def]  | 
|
| 3328 | 286  | 
"!!g. [| kg: nat; \  | 
| 1461 | 287  | 
\ ALL l:list(nat). g`l < ack(kg, list_add(l)); \  | 
288  | 
\         fs : list({f: primrec .                               \
 | 
|
289  | 
\ EX kf:nat. ALL l:list(nat). \  | 
|
290  | 
\ f`l < ack(kf, list_add(l))}) \  | 
|
| 515 | 291  | 
\ |] ==> EX k:nat. ALL l: list(nat). COMP(g,fs)`l < ack(k, list_add(l))";  | 
| 2469 | 292  | 
by (Asm_simp_tac 1);  | 
| 515 | 293  | 
by (forward_tac [list_CollectD] 1);  | 
294  | 
by (etac (COMP_map_lemma RS bexE) 1);  | 
|
295  | 
by (rtac (ballI RS bexI) 1);  | 
|
296  | 
by (etac (bspec RS lt_trans) 1);  | 
|
297  | 
by (rtac lt_trans 2);  | 
|
298  | 
by (rtac ack_nest_bound 3);  | 
|
299  | 
by (etac (bspec RS ack_lt_mono2) 2);  | 
|
300  | 
by (tc_tac [map_type]);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
301  | 
qed "COMP_case";  | 
| 515 | 302  | 
|
303  | 
(** PREC case **)  | 
|
304  | 
||
305  | 
goalw Primrec.thy [PREC_def]  | 
|
| 1461 | 306  | 
"!!f g. [| ALL l:list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \  | 
307  | 
\ ALL l:list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); \  | 
|
308  | 
\ f: primrec; kf: nat; \  | 
|
309  | 
\ g: primrec; kg: nat; \  | 
|
310  | 
\ l: list(nat) \  | 
|
| 515 | 311  | 
\ |] ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))";  | 
312  | 
by (etac list.elim 1);  | 
|
| 2469 | 313  | 
by (asm_simp_tac (!simpset addsimps [[nat_le_refl, lt_ack2] MRS lt_trans]) 1);  | 
314  | 
by (Asm_simp_tac 1);  | 
|
| 515 | 315  | 
by (etac ssubst 1); (*get rid of the needless assumption*)  | 
316  | 
by (eres_inst_tac [("n","a")] nat_induct 1);
 | 
|
317  | 
(*base case*)  | 
|
| 2469 | 318  | 
by (DO_GOAL [Asm_simp_tac, rtac lt_trans, etac bspec,  | 
| 1461 | 319  | 
assume_tac, rtac (add_le_self RS ack_lt_mono1),  | 
320  | 
REPEAT o ares_tac (ack_typechecks)] 1);  | 
|
| 515 | 321  | 
(*ind step*)  | 
| 3328 | 322  | 
by (Asm_simp_tac 1);  | 
| 515 | 323  | 
by (rtac (succ_leI RS lt_trans1) 1);  | 
324  | 
by (res_inst_tac [("j", "g ` ?ll #+ ?mm")] lt_trans1 1);
 | 
|
325  | 
by (etac bspec 2);  | 
|
326  | 
by (rtac (nat_le_refl RS add_le_mono) 1);  | 
|
327  | 
by (tc_tac []);  | 
|
| 2469 | 328  | 
by (asm_simp_tac (!simpset addsimps [add_le_self2]) 1);  | 
| 515 | 329  | 
(*final part of the simplification*)  | 
| 2469 | 330  | 
by (Asm_simp_tac 1);  | 
| 515 | 331  | 
by (rtac (add_le_self2 RS ack_le_mono1 RS lt_trans1) 1);  | 
332  | 
by (etac ack_lt_mono2 5);  | 
|
333  | 
by (tc_tac []);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
334  | 
qed "PREC_case_lemma";  | 
| 515 | 335  | 
|
336  | 
goal Primrec.thy  | 
|
| 1461 | 337  | 
"!!f g. [| f: primrec; kf: nat; \  | 
338  | 
\ g: primrec; kg: nat; \  | 
|
339  | 
\ ALL l:list(nat). f`l < ack(kf, list_add(l)); \  | 
|
340  | 
\ ALL l:list(nat). g`l < ack(kg, list_add(l)) \  | 
|
341  | 
\ |] ==> EX k:nat. ALL l: list(nat). \  | 
|
342  | 
\ PREC(f,g)`l< ack(k, list_add(l))";  | 
|
| 515 | 343  | 
by (rtac (ballI RS bexI) 1);  | 
344  | 
by (rtac ([add_le_self, PREC_case_lemma] MRS lt_trans1) 1);  | 
|
345  | 
by (REPEAT  | 
|
346  | 
(SOMEGOAL  | 
|
347  | 
(FIRST' [test_assume_tac,  | 
|
| 1461 | 348  | 
match_tac (ack_typechecks),  | 
349  | 
rtac (ack_add_bound2 RS ballI) THEN' etac bspec])));  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
350  | 
qed "PREC_case";  | 
| 515 | 351  | 
|
352  | 
goal Primrec.thy  | 
|
353  | 
"!!f. f:primrec ==> EX k:nat. ALL l:list(nat). f`l < ack(k, list_add(l))";  | 
|
354  | 
by (etac primrec.induct 1);  | 
|
| 2469 | 355  | 
by (safe_tac (!claset));  | 
| 515 | 356  | 
by (DEPTH_SOLVE  | 
357  | 
(ares_tac ([SC_case, CONST_case, PROJ_case, COMP_case, PREC_case,  | 
|
| 1461 | 358  | 
bexI, ballI] @ nat_typechecks) 1));  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
359  | 
qed "ack_bounds_primrec";  | 
| 515 | 360  | 
|
361  | 
goal Primrec.thy  | 
|
362  | 
"~ (lam l:list(nat). list_case(0, %x xs. ack(x,x), l)) : primrec";  | 
|
363  | 
by (rtac notI 1);  | 
|
364  | 
by (etac (ack_bounds_primrec RS bexE) 1);  | 
|
365  | 
by (rtac lt_irrefl 1);  | 
|
366  | 
by (dres_inst_tac [("x", "[x]")] bspec 1);
 | 
|
| 2469 | 367  | 
by (Asm_simp_tac 1);  | 
368  | 
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
 | 
369  | 
qed "ack_not_primrec";  | 
| 515 | 370  |