| author | paulson | 
| Thu, 27 Sep 2007 17:55:28 +0200 | |
| changeset 24742 | 73b8b42a36b6 | 
| parent 24551 | af7ef6bcc149 | 
| child 25157 | 8b80535cd017 | 
| permissions | -rw-r--r-- | 
| 11024 | 1  | 
(* Title: HOL/ex/Primrec.thy  | 
| 3335 | 2  | 
ID: $Id$  | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1997 University of Cambridge  | 
|
5  | 
||
| 11024 | 6  | 
Primitive Recursive Functions. Demonstrates recursive definitions,  | 
7  | 
the TFL package.  | 
|
| 3335 | 8  | 
*)  | 
9  | 
||
| 11024 | 10  | 
header {* Primitive Recursive Functions *}
 | 
11  | 
||
| 16417 | 12  | 
theory Primrec imports Main begin  | 
| 11024 | 13  | 
|
14  | 
text {*
 | 
|
15  | 
Proof adopted from  | 
|
16  | 
||
17  | 
Nora Szasz, A Machine Checked Proof that Ackermann's Function is not  | 
|
18  | 
Primitive Recursive, In: Huet \& Plotkin, eds., Logical Environments  | 
|
19  | 
(CUP, 1993), 317-338.  | 
|
20  | 
||
21  | 
See also E. Mendelson, Introduction to Mathematical Logic. (Van  | 
|
22  | 
Nostrand, 1964), page 250, exercise 11.  | 
|
23  | 
\medskip  | 
|
24  | 
*}  | 
|
25  | 
||
26  | 
consts ack :: "nat * nat => nat"  | 
|
27  | 
recdef ack "less_than <*lex*> less_than"  | 
|
28  | 
"ack (0, n) = Suc n"  | 
|
29  | 
"ack (Suc m, 0) = ack (m, 1)"  | 
|
30  | 
"ack (Suc m, Suc n) = ack (m, ack (Suc m, n))"  | 
|
31  | 
||
32  | 
consts list_add :: "nat list => nat"  | 
|
33  | 
primrec  | 
|
34  | 
"list_add [] = 0"  | 
|
35  | 
"list_add (m # ms) = m + list_add ms"  | 
|
36  | 
||
37  | 
consts zeroHd :: "nat list => nat"  | 
|
38  | 
primrec  | 
|
39  | 
"zeroHd [] = 0"  | 
|
40  | 
"zeroHd (m # ms) = m"  | 
|
41  | 
||
42  | 
||
43  | 
text {* The set of primitive recursive functions of type @{typ "nat list => nat"}. *}
 | 
|
44  | 
||
| 19736 | 45  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
46  | 
SC :: "nat list => nat" where  | 
| 19736 | 47  | 
"SC l = Suc (zeroHd l)"  | 
| 3335 | 48  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
49  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
50  | 
CONSTANT :: "nat => nat list => nat" where  | 
| 19736 | 51  | 
"CONSTANT k l = k"  | 
| 11024 | 52  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
53  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
54  | 
PROJ :: "nat => nat list => nat" where  | 
| 19736 | 55  | 
"PROJ i l = zeroHd (drop i l)"  | 
| 11024 | 56  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
57  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
58  | 
COMP :: "(nat list => nat) => (nat list => nat) list => nat list => nat" where  | 
| 19736 | 59  | 
"COMP g fs l = g (map (\<lambda>f. f l) fs)"  | 
| 11024 | 60  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
61  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
62  | 
PREC :: "(nat list => nat) => (nat list => nat) => nat list => nat" where  | 
| 19736 | 63  | 
"PREC f g l =  | 
64  | 
(case l of  | 
|
| 11024 | 65  | 
[] => 0  | 
| 19736 | 66  | 
| x # l' => nat_rec (f l') (\<lambda>y r. g (r # y # l')) x)"  | 
| 11024 | 67  | 
  -- {* Note that @{term g} is applied first to @{term "PREC f g y"} and then to @{term y}! *}
 | 
68  | 
||
| 23776 | 69  | 
inductive PRIMREC :: "(nat list => nat) => bool"  | 
| 22283 | 70  | 
where  | 
71  | 
SC: "PRIMREC SC"  | 
|
72  | 
| CONSTANT: "PRIMREC (CONSTANT k)"  | 
|
73  | 
| PROJ: "PRIMREC (PROJ i)"  | 
|
| 22944 | 74  | 
| COMP: "PRIMREC g ==> \<forall>f \<in> set fs. PRIMREC f ==> PRIMREC (COMP g fs)"  | 
| 22283 | 75  | 
| PREC: "PRIMREC f ==> PRIMREC g ==> PRIMREC (PREC f g)"  | 
| 11024 | 76  | 
|
77  | 
||
78  | 
text {* Useful special cases of evaluation *}
 | 
|
79  | 
||
80  | 
lemma SC [simp]: "SC (x # l) = Suc x"  | 
|
81  | 
apply (simp add: SC_def)  | 
|
82  | 
done  | 
|
83  | 
||
| 19676 | 84  | 
lemma CONSTANT [simp]: "CONSTANT k l = k"  | 
85  | 
apply (simp add: CONSTANT_def)  | 
|
| 11024 | 86  | 
done  | 
87  | 
||
88  | 
lemma PROJ_0 [simp]: "PROJ 0 (x # l) = x"  | 
|
89  | 
apply (simp add: PROJ_def)  | 
|
90  | 
done  | 
|
91  | 
||
92  | 
lemma COMP_1 [simp]: "COMP g [f] l = g [f l]"  | 
|
93  | 
apply (simp add: COMP_def)  | 
|
94  | 
done  | 
|
| 3335 | 95  | 
|
| 11024 | 96  | 
lemma PREC_0 [simp]: "PREC f g (0 # l) = f l"  | 
97  | 
apply (simp add: PREC_def)  | 
|
98  | 
done  | 
|
99  | 
||
100  | 
lemma PREC_Suc [simp]: "PREC f g (Suc x # l) = g (PREC f g (x # l) # x # l)"  | 
|
101  | 
apply (simp add: PREC_def)  | 
|
102  | 
done  | 
|
103  | 
||
104  | 
||
105  | 
text {* PROPERTY A 4 *}
 | 
|
106  | 
||
107  | 
lemma less_ack2 [iff]: "j < ack (i, j)"  | 
|
108  | 
apply (induct i j rule: ack.induct)  | 
|
109  | 
apply simp_all  | 
|
110  | 
done  | 
|
111  | 
||
112  | 
||
113  | 
text {* PROPERTY A 5-, the single-step lemma *}
 | 
|
114  | 
||
115  | 
lemma ack_less_ack_Suc2 [iff]: "ack(i, j) < ack (i, Suc j)"  | 
|
116  | 
apply (induct i j rule: ack.induct)  | 
|
117  | 
apply simp_all  | 
|
118  | 
done  | 
|
119  | 
||
120  | 
||
121  | 
text {* PROPERTY A 5, monotonicity for @{text "<"} *}
 | 
|
122  | 
||
123  | 
lemma ack_less_mono2: "j < k ==> ack (i, j) < ack (i, k)"  | 
|
124  | 
apply (induct i k rule: ack.induct)  | 
|
125  | 
apply simp_all  | 
|
126  | 
apply (blast elim!: less_SucE intro: less_trans)  | 
|
127  | 
done  | 
|
128  | 
||
129  | 
||
130  | 
text {* PROPERTY A 5', monotonicity for @{text \<le>} *}
 | 
|
131  | 
||
132  | 
lemma ack_le_mono2: "j \<le> k ==> ack (i, j) \<le> ack (i, k)"  | 
|
133  | 
apply (simp add: order_le_less)  | 
|
134  | 
apply (blast intro: ack_less_mono2)  | 
|
135  | 
done  | 
|
| 3335 | 136  | 
|
| 11024 | 137  | 
|
138  | 
text {* PROPERTY A 6 *}
 | 
|
139  | 
||
140  | 
lemma ack2_le_ack1 [iff]: "ack (i, Suc j) \<le> ack (Suc i, j)"  | 
|
141  | 
apply (induct j)  | 
|
142  | 
apply simp_all  | 
|
143  | 
apply (blast intro: ack_le_mono2 less_ack2 [THEN Suc_leI] le_trans)  | 
|
144  | 
done  | 
|
145  | 
||
146  | 
||
147  | 
text {* PROPERTY A 7-, the single-step lemma *}
 | 
|
148  | 
||
149  | 
lemma ack_less_ack_Suc1 [iff]: "ack (i, j) < ack (Suc i, j)"  | 
|
150  | 
apply (blast intro: ack_less_mono2 less_le_trans)  | 
|
151  | 
done  | 
|
152  | 
||
153  | 
||
| 19676 | 154  | 
text {* PROPERTY A 4'? Extra lemma needed for @{term CONSTANT} case, constant functions *}
 | 
| 11024 | 155  | 
|
156  | 
lemma less_ack1 [iff]: "i < ack (i, j)"  | 
|
157  | 
apply (induct i)  | 
|
158  | 
apply simp_all  | 
|
159  | 
apply (blast intro: Suc_leI le_less_trans)  | 
|
160  | 
done  | 
|
161  | 
||
162  | 
||
163  | 
text {* PROPERTY A 8 *}
 | 
|
164  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
165  | 
lemma ack_1 [simp]: "ack (Suc 0, j) = j + 2"  | 
| 11024 | 166  | 
apply (induct j)  | 
167  | 
apply simp_all  | 
|
168  | 
done  | 
|
169  | 
||
170  | 
||
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
171  | 
text {* PROPERTY A 9.  The unary @{text 1} and @{text 2} in @{term
 | 
| 11024 | 172  | 
ack} is essential for the rewriting. *}  | 
173  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
174  | 
lemma ack_2 [simp]: "ack (Suc (Suc 0), j) = 2 * j + 3"  | 
| 11024 | 175  | 
apply (induct j)  | 
176  | 
apply simp_all  | 
|
177  | 
done  | 
|
| 3335 | 178  | 
|
179  | 
||
| 11024 | 180  | 
text {* PROPERTY A 7, monotonicity for @{text "<"} [not clear why
 | 
181  | 
  @{thm [source] ack_1} is now needed first!] *}
 | 
|
182  | 
||
183  | 
lemma ack_less_mono1_aux: "ack (i, k) < ack (Suc (i +i'), k)"  | 
|
184  | 
apply (induct i k rule: ack.induct)  | 
|
185  | 
apply simp_all  | 
|
186  | 
prefer 2  | 
|
187  | 
apply (blast intro: less_trans ack_less_mono2)  | 
|
188  | 
apply (induct_tac i' n rule: ack.induct)  | 
|
189  | 
apply simp_all  | 
|
190  | 
apply (blast intro: Suc_leI [THEN le_less_trans] ack_less_mono2)  | 
|
191  | 
done  | 
|
192  | 
||
193  | 
lemma ack_less_mono1: "i < j ==> ack (i, k) < ack (j, k)"  | 
|
194  | 
apply (drule less_imp_Suc_add)  | 
|
195  | 
apply (blast intro!: ack_less_mono1_aux)  | 
|
196  | 
done  | 
|
197  | 
||
198  | 
||
199  | 
text {* PROPERTY A 7', monotonicity for @{text "\<le>"} *}
 | 
|
200  | 
||
201  | 
lemma ack_le_mono1: "i \<le> j ==> ack (i, k) \<le> ack (j, k)"  | 
|
202  | 
apply (simp add: order_le_less)  | 
|
203  | 
apply (blast intro: ack_less_mono1)  | 
|
204  | 
done  | 
|
205  | 
||
206  | 
||
207  | 
text {* PROPERTY A 10 *}
 | 
|
208  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
209  | 
lemma ack_nest_bound: "ack(i1, ack (i2, j)) < ack (2 + (i1 + i2), j)"  | 
| 11024 | 210  | 
apply (simp add: numerals)  | 
211  | 
apply (rule ack2_le_ack1 [THEN [2] less_le_trans])  | 
|
212  | 
apply simp  | 
|
213  | 
apply (rule le_add1 [THEN ack_le_mono1, THEN le_less_trans])  | 
|
214  | 
apply (rule ack_less_mono1 [THEN ack_less_mono2])  | 
|
215  | 
apply (simp add: le_imp_less_Suc le_add2)  | 
|
216  | 
done  | 
|
217  | 
||
| 3335 | 218  | 
|
| 11024 | 219  | 
text {* PROPERTY A 11 *}
 | 
| 3335 | 220  | 
|
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
221  | 
lemma ack_add_bound: "ack (i1, j) + ack (i2, j) < ack (4 + (i1 + i2), j)"  | 
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24551 
diff
changeset
 | 
222  | 
apply (rule less_trans [of _ "ack (Suc (Suc 0), ack (i1 + i2, j))" _])  | 
| 11024 | 223  | 
prefer 2  | 
224  | 
apply (rule ack_nest_bound [THEN less_le_trans])  | 
|
225  | 
apply (simp add: Suc3_eq_add_3)  | 
|
226  | 
apply simp  | 
|
227  | 
apply (cut_tac i = i1 and m1 = i2 and k = j in le_add1 [THEN ack_le_mono1])  | 
|
228  | 
apply (cut_tac i = "i2" and m1 = i1 and k = j in le_add2 [THEN ack_le_mono1])  | 
|
229  | 
apply auto  | 
|
230  | 
done  | 
|
231  | 
||
232  | 
||
233  | 
text {* PROPERTY A 12.  Article uses existential quantifier but the ALF proof
 | 
|
234  | 
  used @{text "k + 4"}.  Quantified version must be nested @{text
 | 
|
235  | 
"\<exists>k'. \<forall>i j. ..."} *}  | 
|
| 3335 | 236  | 
|
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
237  | 
lemma ack_add_bound2: "i < ack (k, j) ==> i + j < ack (4 + k, j)"  | 
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24551 
diff
changeset
 | 
238  | 
apply (rule less_trans [of _ "ack (k, j) + ack (0, j)" _])  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24551 
diff
changeset
 | 
239  | 
apply (blast intro: add_less_mono less_ack2)  | 
| 11024 | 240  | 
apply (rule ack_add_bound [THEN less_le_trans])  | 
241  | 
apply simp  | 
|
242  | 
done  | 
|
243  | 
||
244  | 
||
245  | 
||
246  | 
text {* Inductive definition of the @{term PR} functions *}
 | 
|
| 3335 | 247  | 
|
| 11024 | 248  | 
text {* MAIN RESULT *}
 | 
249  | 
||
250  | 
lemma SC_case: "SC l < ack (1, list_add l)"  | 
|
251  | 
apply (unfold SC_def)  | 
|
252  | 
apply (induct l)  | 
|
253  | 
apply (simp_all add: le_add1 le_imp_less_Suc)  | 
|
254  | 
done  | 
|
255  | 
||
| 19676 | 256  | 
lemma CONSTANT_case: "CONSTANT k l < ack (k, list_add l)"  | 
| 24551 | 257  | 
by simp  | 
| 3335 | 258  | 
|
| 11024 | 259  | 
lemma PROJ_case [rule_format]: "\<forall>i. PROJ i l < ack (0, list_add l)"  | 
260  | 
apply (simp add: PROJ_def)  | 
|
261  | 
apply (induct l)  | 
|
| 24551 | 262  | 
apply (auto simp add: drop_Cons split: nat.split)  | 
263  | 
apply (blast intro: less_le_trans le_add2)  | 
|
| 11024 | 264  | 
done  | 
265  | 
||
266  | 
||
267  | 
text {* @{term COMP} case *}
 | 
|
| 3335 | 268  | 
|
| 22944 | 269  | 
lemma COMP_map_aux: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack (kf, list_add l))  | 
| 11024 | 270  | 
==> \<exists>k. \<forall>l. list_add (map (\<lambda>f. f l) fs) < ack (k, list_add l)"  | 
| 22944 | 271  | 
apply (induct fs)  | 
| 24551 | 272  | 
apply (rule_tac x = 0 in exI)  | 
| 11024 | 273  | 
apply simp  | 
274  | 
apply simp  | 
|
275  | 
apply (blast intro: add_less_mono ack_add_bound less_trans)  | 
|
276  | 
done  | 
|
277  | 
||
278  | 
lemma COMP_case:  | 
|
279  | 
"\<forall>l. g l < ack (kg, list_add l) ==>  | 
|
| 22944 | 280  | 
\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack(kf, list_add l))  | 
| 11024 | 281  | 
==> \<exists>k. \<forall>l. COMP g fs l < ack(k, list_add l)"  | 
282  | 
apply (unfold COMP_def)  | 
|
| 16588 | 283  | 
    --{*Now, if meson tolerated map, we could finish with
 | 
| 16731 | 284  | 
  @{text "(drule COMP_map_aux, meson ack_less_mono2 ack_nest_bound less_trans)"} *}
 | 
| 16588 | 285  | 
apply (erule COMP_map_aux [THEN exE])  | 
286  | 
apply (rule exI)  | 
|
287  | 
apply (rule allI)  | 
|
288  | 
apply (drule spec)+  | 
|
289  | 
apply (erule less_trans)  | 
|
290  | 
apply (blast intro: ack_less_mono2 ack_nest_bound less_trans)  | 
|
| 11024 | 291  | 
done  | 
292  | 
||
293  | 
||
294  | 
text {* @{term PREC} case *}
 | 
|
| 3335 | 295  | 
|
| 11024 | 296  | 
lemma PREC_case_aux:  | 
297  | 
"\<forall>l. f l + list_add l < ack (kf, list_add l) ==>  | 
|
298  | 
\<forall>l. g l + list_add l < ack (kg, list_add l) ==>  | 
|
299  | 
PREC f g l + list_add l < ack (Suc (kf + kg), list_add l)"  | 
|
300  | 
apply (unfold PREC_def)  | 
|
301  | 
apply (case_tac l)  | 
|
302  | 
apply simp_all  | 
|
303  | 
apply (blast intro: less_trans)  | 
|
304  | 
  apply (erule ssubst) -- {* get rid of the needless assumption *}
 | 
|
305  | 
apply (induct_tac a)  | 
|
306  | 
apply simp_all  | 
|
307  | 
   txt {* base case *}
 | 
|
308  | 
apply (blast intro: le_add1 [THEN le_imp_less_Suc, THEN ack_less_mono1] less_trans)  | 
|
309  | 
  txt {* induction step *}
 | 
|
310  | 
apply (rule Suc_leI [THEN le_less_trans])  | 
|
311  | 
apply (rule le_refl [THEN add_le_mono, THEN le_less_trans])  | 
|
312  | 
prefer 2  | 
|
313  | 
apply (erule spec)  | 
|
314  | 
apply (simp add: le_add2)  | 
|
315  | 
  txt {* final part of the simplification *}
 | 
|
316  | 
apply simp  | 
|
317  | 
apply (rule le_add2 [THEN ack_le_mono1, THEN le_less_trans])  | 
|
318  | 
apply (erule ack_less_mono2)  | 
|
319  | 
done  | 
|
320  | 
||
321  | 
lemma PREC_case:  | 
|
322  | 
"\<forall>l. f l < ack (kf, list_add l) ==>  | 
|
323  | 
\<forall>l. g l < ack (kg, list_add l) ==>  | 
|
324  | 
\<exists>k. \<forall>l. PREC f g l < ack (k, list_add l)"  | 
|
| 24551 | 325  | 
by (metis le_less_trans [OF le_add1 PREC_case_aux] ack_add_bound2)  | 
| 11024 | 326  | 
|
| 22283 | 327  | 
lemma ack_bounds_PRIMREC: "PRIMREC f ==> \<exists>k. \<forall>l. f l < ack (k, list_add l)"  | 
| 11024 | 328  | 
apply (erule PRIMREC.induct)  | 
| 19676 | 329  | 
apply (blast intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)+  | 
| 11024 | 330  | 
done  | 
331  | 
||
| 22283 | 332  | 
lemma ack_not_PRIMREC: "\<not> PRIMREC (\<lambda>l. case l of [] => 0 | x # l' => ack (x, x))"  | 
| 11024 | 333  | 
apply (rule notI)  | 
334  | 
apply (erule ack_bounds_PRIMREC [THEN exE])  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24551 
diff
changeset
 | 
335  | 
apply (rule Nat.less_irrefl)  | 
| 11024 | 336  | 
apply (drule_tac x = "[x]" in spec)  | 
337  | 
apply simp  | 
|
338  | 
done  | 
|
| 3335 | 339  | 
|
340  | 
end  |