author | paulson |
Thu, 12 Sep 2019 14:51:50 +0100 | |
changeset 70689 | 67360d50ebb3 |
parent 69880 | fe3c12990893 |
child 75014 | 0778d233964d |
permissions | -rw-r--r-- |
11024 | 1 |
(* Title: HOL/ex/Primrec.thy |
3335 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1997 University of Cambridge |
|
4 |
*) |
|
5 |
||
69880 | 6 |
section \<open>Ackermann's Function and the Primitive Recursive Functions\<close> |
11024 | 7 |
|
16417 | 8 |
theory Primrec imports Main begin |
11024 | 9 |
|
61343 | 10 |
text \<open> |
11024 | 11 |
Proof adopted from |
12 |
||
13 |
Nora Szasz, A Machine Checked Proof that Ackermann's Function is not |
|
14 |
Primitive Recursive, In: Huet \& Plotkin, eds., Logical Environments |
|
15 |
(CUP, 1993), 317-338. |
|
16 |
||
17 |
See also E. Mendelson, Introduction to Mathematical Logic. (Van |
|
18 |
Nostrand, 1964), page 250, exercise 11. |
|
19 |
\medskip |
|
61343 | 20 |
\<close> |
11024 | 21 |
|
22 |
||
61343 | 23 |
subsection\<open>Ackermann's Function\<close> |
11024 | 24 |
|
69880 | 25 |
fun ack :: "[nat,nat] \<Rightarrow> nat" where |
26 |
"ack 0 n = Suc n" |
|
27 |
| "ack (Suc m) 0 = ack m 1" |
|
28 |
| "ack (Suc m) (Suc n) = ack m (ack (Suc m) n)" |
|
11024 | 29 |
|
30 |
||
61343 | 31 |
text \<open>PROPERTY A 4\<close> |
11024 | 32 |
|
27626 | 33 |
lemma less_ack2 [iff]: "j < ack i j" |
69880 | 34 |
by (induct i j rule: ack.induct) simp_all |
11024 | 35 |
|
36 |
||
61343 | 37 |
text \<open>PROPERTY A 5-, the single-step lemma\<close> |
11024 | 38 |
|
27626 | 39 |
lemma ack_less_ack_Suc2 [iff]: "ack i j < ack i (Suc j)" |
69880 | 40 |
by (induct i j rule: ack.induct) simp_all |
11024 | 41 |
|
42 |
||
61933 | 43 |
text \<open>PROPERTY A 5, monotonicity for \<open><\<close>\<close> |
11024 | 44 |
|
69880 | 45 |
lemma ack_less_mono2: "j < k \<Longrightarrow> ack i j < ack i k" |
46 |
by (simp add: lift_Suc_mono_less) |
|
11024 | 47 |
|
48 |
||
61933 | 49 |
text \<open>PROPERTY A 5', monotonicity for \<open>\<le>\<close>\<close> |
11024 | 50 |
|
69880 | 51 |
lemma ack_le_mono2: "j \<le> k \<Longrightarrow> ack i j \<le> ack i k" |
52 |
by (simp add: ack_less_mono2 less_mono_imp_le_mono) |
|
3335 | 53 |
|
11024 | 54 |
|
61343 | 55 |
text \<open>PROPERTY A 6\<close> |
11024 | 56 |
|
27626 | 57 |
lemma ack2_le_ack1 [iff]: "ack i (Suc j) \<le> ack (Suc i) j" |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
25157
diff
changeset
|
58 |
proof (induct j) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
25157
diff
changeset
|
59 |
case 0 show ?case by simp |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
25157
diff
changeset
|
60 |
next |
69880 | 61 |
case (Suc j) show ?case |
62 |
by (metis Suc ack.simps(3) ack_le_mono2 le_trans less_ack2 less_eq_Suc_le) |
|
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
25157
diff
changeset
|
63 |
qed |
11024 | 64 |
|
65 |
||
61343 | 66 |
text \<open>PROPERTY A 7-, the single-step lemma\<close> |
11024 | 67 |
|
27626 | 68 |
lemma ack_less_ack_Suc1 [iff]: "ack i j < ack (Suc i) j" |
69880 | 69 |
by (blast intro: ack_less_mono2 less_le_trans) |
11024 | 70 |
|
71 |
||
69597 | 72 |
text \<open>PROPERTY A 4'? Extra lemma needed for \<^term>\<open>CONSTANT\<close> case, constant functions\<close> |
11024 | 73 |
|
27626 | 74 |
lemma less_ack1 [iff]: "i < ack i j" |
69880 | 75 |
proof (induct i) |
76 |
case 0 |
|
77 |
then show ?case |
|
78 |
by simp |
|
79 |
next |
|
80 |
case (Suc i) |
|
81 |
then show ?case |
|
82 |
using less_trans_Suc by blast |
|
83 |
qed |
|
11024 | 84 |
|
85 |
||
61343 | 86 |
text \<open>PROPERTY A 8\<close> |
11024 | 87 |
|
27626 | 88 |
lemma ack_1 [simp]: "ack (Suc 0) j = j + 2" |
69880 | 89 |
by (induct j) simp_all |
11024 | 90 |
|
91 |
||
69597 | 92 |
text \<open>PROPERTY A 9. The unary \<open>1\<close> and \<open>2\<close> in \<^term>\<open>ack\<close> is essential for the rewriting.\<close> |
11024 | 93 |
|
27626 | 94 |
lemma ack_2 [simp]: "ack (Suc (Suc 0)) j = 2 * j + 3" |
69880 | 95 |
by (induct j) simp_all |
3335 | 96 |
|
97 |
||
61933 | 98 |
text \<open>PROPERTY A 7, monotonicity for \<open><\<close> [not clear why |
61343 | 99 |
@{thm [source] ack_1} is now needed first!]\<close> |
11024 | 100 |
|
27626 | 101 |
lemma ack_less_mono1_aux: "ack i k < ack (Suc (i +i')) k" |
34055 | 102 |
proof (induct i k rule: ack.induct) |
103 |
case (1 n) show ?case |
|
69880 | 104 |
using less_le_trans by auto |
34055 | 105 |
next |
106 |
case (2 m) thus ?case by simp |
|
107 |
next |
|
108 |
case (3 m n) thus ?case |
|
69880 | 109 |
using ack_less_mono2 less_trans by fastforce |
34055 | 110 |
qed |
11024 | 111 |
|
69880 | 112 |
lemma ack_less_mono1: "i < j \<Longrightarrow> ack i k < ack j k" |
113 |
using ack_less_mono1_aux less_iff_Suc_add by auto |
|
11024 | 114 |
|
115 |
||
61933 | 116 |
text \<open>PROPERTY A 7', monotonicity for \<open>\<le>\<close>\<close> |
11024 | 117 |
|
69880 | 118 |
lemma ack_le_mono1: "i \<le> j \<Longrightarrow> ack i k \<le> ack j k" |
119 |
using ack_less_mono1 le_eq_less_or_eq by auto |
|
11024 | 120 |
|
121 |
||
61343 | 122 |
text \<open>PROPERTY A 10\<close> |
11024 | 123 |
|
27626 | 124 |
lemma ack_nest_bound: "ack i1 (ack i2 j) < ack (2 + (i1 + i2)) j" |
69880 | 125 |
proof - |
126 |
have "ack i1 (ack i2 j) < ack (i1 + i2) (ack (Suc (i1 + i2)) j)" |
|
127 |
by (meson ack_le_mono1 ack_less_mono1 ack_less_mono2 le_add1 le_trans less_add_Suc2 not_less) |
|
128 |
also have "... = ack (Suc (i1 + i2)) (Suc j)" |
|
129 |
by simp |
|
130 |
also have "... \<le> ack (2 + (i1 + i2)) j" |
|
131 |
using ack2_le_ack1 add_2_eq_Suc by presburger |
|
132 |
finally show ?thesis . |
|
133 |
qed |
|
134 |
||
11024 | 135 |
|
3335 | 136 |
|
61343 | 137 |
text \<open>PROPERTY A 11\<close> |
3335 | 138 |
|
27626 | 139 |
lemma ack_add_bound: "ack i1 j + ack i2 j < ack (4 + (i1 + i2)) j" |
69880 | 140 |
proof - |
141 |
have "ack i1 j \<le> ack (i1 + i2) j" "ack i2 j \<le> ack (i1 + i2) j" |
|
142 |
by (simp_all add: ack_le_mono1) |
|
143 |
then have "ack i1 j + ack i2 j < ack (Suc (Suc 0)) (ack (i1 + i2) j)" |
|
144 |
by simp |
|
145 |
also have "... < ack (4 + (i1 + i2)) j" |
|
146 |
by (metis ack_nest_bound add.assoc numeral_2_eq_2 numeral_Bit0) |
|
147 |
finally show ?thesis . |
|
148 |
qed |
|
11024 | 149 |
|
150 |
||
61343 | 151 |
text \<open>PROPERTY A 12. Article uses existential quantifier but the ALF proof |
61933 | 152 |
used \<open>k + 4\<close>. Quantified version must be nested \<open>\<exists>k'. \<forall>i j. ...\<close>\<close> |
3335 | 153 |
|
69880 | 154 |
lemma ack_add_bound2: |
155 |
assumes "i < ack k j" shows "i + j < ack (4 + k) j" |
|
156 |
proof - |
|
157 |
have "i + j < ack k j + ack 0 j" |
|
158 |
using assms by auto |
|
159 |
also have "... < ack (4 + k) j" |
|
160 |
by (metis ack_add_bound add.right_neutral) |
|
161 |
finally show ?thesis . |
|
162 |
qed |
|
27626 | 163 |
|
164 |
||
61343 | 165 |
subsection\<open>Primitive Recursive Functions\<close> |
27626 | 166 |
|
69880 | 167 |
primrec hd0 :: "nat list \<Rightarrow> nat" where |
168 |
"hd0 [] = 0" |
|
169 |
| "hd0 (m # ms) = m" |
|
11024 | 170 |
|
171 |
||
69880 | 172 |
text \<open>Inductive definition of the set of primitive recursive functions of type \<^typ>\<open>nat list \<Rightarrow> nat\<close>.\<close> |
11024 | 173 |
|
69880 | 174 |
definition SC :: "nat list \<Rightarrow> nat" |
175 |
where "SC l = Suc (hd0 l)" |
|
27626 | 176 |
|
69880 | 177 |
definition CONSTANT :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
178 |
where "CONSTANT k l = k" |
|
27626 | 179 |
|
69880 | 180 |
definition PROJ :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
181 |
where "PROJ i l = hd0 (drop i l)" |
|
27626 | 182 |
|
69880 | 183 |
definition COMP :: "[nat list \<Rightarrow> nat, (nat list \<Rightarrow> nat) list, nat list] \<Rightarrow> nat" |
184 |
where "COMP g fs l = g (map (\<lambda>f. f l) fs)" |
|
27626 | 185 |
|
69880 | 186 |
fun PREC :: "[nat list \<Rightarrow> nat, nat list \<Rightarrow> nat, nat list] \<Rightarrow> nat" |
187 |
where |
|
188 |
"PREC f g [] = 0" |
|
189 |
| "PREC f g (x # l) = rec_nat (f l) (\<lambda>y r. g (r # y # l)) x" |
|
190 |
\<comment> \<open>Note that \<^term>\<open>g\<close> is applied first to \<^term>\<open>PREC f g y\<close> and then to \<^term>\<open>y\<close>!\<close> |
|
191 |
||
192 |
inductive PRIMREC :: "(nat list \<Rightarrow> nat) \<Rightarrow> bool" where |
|
193 |
SC: "PRIMREC SC" |
|
194 |
| CONSTANT: "PRIMREC (CONSTANT k)" |
|
195 |
| PROJ: "PRIMREC (PROJ i)" |
|
196 |
| COMP: "PRIMREC g \<Longrightarrow> \<forall>f \<in> set fs. PRIMREC f \<Longrightarrow> PRIMREC (COMP g fs)" |
|
197 |
| PREC: "PRIMREC f \<Longrightarrow> PRIMREC g \<Longrightarrow> PRIMREC (PREC f g)" |
|
27626 | 198 |
|
199 |
||
61343 | 200 |
text \<open>Useful special cases of evaluation\<close> |
27626 | 201 |
|
202 |
lemma SC [simp]: "SC (x # l) = Suc x" |
|
69880 | 203 |
by (simp add: SC_def) |
27626 | 204 |
|
205 |
lemma PROJ_0 [simp]: "PROJ 0 (x # l) = x" |
|
69880 | 206 |
by (simp add: PROJ_def) |
27626 | 207 |
|
208 |
lemma COMP_1 [simp]: "COMP g [f] l = g [f l]" |
|
69880 | 209 |
by (simp add: COMP_def) |
27626 | 210 |
|
69880 | 211 |
lemma PREC_0: "PREC f g (0 # l) = f l" |
212 |
by simp |
|
27626 | 213 |
|
214 |
lemma PREC_Suc [simp]: "PREC f g (Suc x # l) = g (PREC f g (x # l) # x # l)" |
|
69880 | 215 |
by auto |
27626 | 216 |
|
3335 | 217 |
|
69880 | 218 |
subsection \<open>MAIN RESULT\<close> |
11024 | 219 |
|
63882
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents:
61933
diff
changeset
|
220 |
lemma SC_case: "SC l < ack 1 (sum_list l)" |
69880 | 221 |
unfolding SC_def |
222 |
by (induct l) (simp_all add: le_add1 le_imp_less_Suc) |
|
11024 | 223 |
|
63882
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents:
61933
diff
changeset
|
224 |
lemma CONSTANT_case: "CONSTANT k l < ack k (sum_list l)" |
69880 | 225 |
by (simp add: CONSTANT_def) |
3335 | 226 |
|
63882
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents:
61933
diff
changeset
|
227 |
lemma PROJ_case: "PROJ i l < ack 0 (sum_list l)" |
69880 | 228 |
unfolding PROJ_def |
229 |
proof (induct l arbitrary: i) |
|
230 |
case Nil |
|
231 |
then show ?case |
|
232 |
by simp |
|
233 |
next |
|
234 |
case (Cons a l) |
|
235 |
then show ?case |
|
236 |
by (metis ack.simps(1) add.commute drop_Cons' hd0.simps(2) leD leI lessI not_less_eq sum_list.Cons trans_le_add2) |
|
237 |
qed |
|
11024 | 238 |
|
239 |
||
69597 | 240 |
text \<open>\<^term>\<open>COMP\<close> case\<close> |
3335 | 241 |
|
63882
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents:
61933
diff
changeset
|
242 |
lemma COMP_map_aux: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (sum_list l)) |
69880 | 243 |
\<Longrightarrow> \<exists>k. \<forall>l. sum_list (map (\<lambda>f. f l) fs) < ack k (sum_list l)" |
244 |
proof (induct fs) |
|
245 |
case Nil |
|
246 |
then show ?case |
|
247 |
by auto |
|
248 |
next |
|
249 |
case (Cons a fs) |
|
250 |
then show ?case |
|
251 |
by simp (blast intro: add_less_mono ack_add_bound less_trans) |
|
252 |
qed |
|
11024 | 253 |
|
254 |
lemma COMP_case: |
|
69880 | 255 |
assumes 1: "\<forall>l. g l < ack kg (sum_list l)" |
256 |
and 2: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (sum_list l))" |
|
257 |
shows "\<exists>k. \<forall>l. COMP g fs l < ack k (sum_list l)" |
|
258 |
unfolding COMP_def |
|
259 |
using 1 COMP_map_aux [OF 2] by (meson ack_less_mono2 ack_nest_bound less_trans) |
|
11024 | 260 |
|
69597 | 261 |
text \<open>\<^term>\<open>PREC\<close> case\<close> |
3335 | 262 |
|
11024 | 263 |
lemma PREC_case_aux: |
69880 | 264 |
assumes f: "\<And>l. f l + sum_list l < ack kf (sum_list l)" |
265 |
and g: "\<And>l. g l + sum_list l < ack kg (sum_list l)" |
|
266 |
shows "PREC f g l + sum_list l < ack (Suc (kf + kg)) (sum_list l)" |
|
267 |
proof (cases l) |
|
268 |
case Nil |
|
269 |
then show ?thesis |
|
270 |
by (simp add: Suc_lessD) |
|
271 |
next |
|
272 |
case (Cons m l) |
|
273 |
have "rec_nat (f l) (\<lambda>y r. g (r # y # l)) m + (m + sum_list l) < ack (Suc (kf + kg)) (m + sum_list l)" |
|
274 |
proof (induct m) |
|
275 |
case 0 |
|
276 |
then show ?case |
|
277 |
using ack_less_mono1_aux f less_trans by fastforce |
|
278 |
next |
|
279 |
case (Suc m) |
|
280 |
let ?r = "rec_nat (f l) (\<lambda>y r. g (r # y # l)) m" |
|
281 |
have "\<not> g (?r # m # l) + sum_list (?r # m # l) < g (?r # m # l) + (m + sum_list l)" |
|
282 |
by force |
|
283 |
then have "g (?r # m # l) + (m + sum_list l) < ack kg (sum_list (?r # m # l))" |
|
284 |
by (meson assms(2) leI less_le_trans) |
|
285 |
moreover |
|
286 |
have "... < ack (kf + kg) (ack (Suc (kf + kg)) (m + sum_list l))" |
|
287 |
using Suc.hyps by simp (meson ack_le_mono1 ack_less_mono2 le_add2 le_less_trans) |
|
288 |
ultimately show ?case |
|
289 |
by auto |
|
290 |
qed |
|
291 |
then show ?thesis |
|
292 |
by (simp add: local.Cons) |
|
293 |
qed |
|
11024 | 294 |
|
69880 | 295 |
proposition PREC_case: |
296 |
"\<lbrakk>\<And>l. f l < ack kf (sum_list l); \<And>l. g l < ack kg (sum_list l)\<rbrakk> |
|
297 |
\<Longrightarrow> \<exists>k. \<forall>l. PREC f g l < ack k (sum_list l)" |
|
298 |
by (metis le_less_trans [OF le_add1 PREC_case_aux] ack_add_bound2) |
|
11024 | 299 |
|
69880 | 300 |
lemma ack_bounds_PRIMREC: "PRIMREC f \<Longrightarrow> \<exists>k. \<forall>l. f l < ack k (sum_list l)" |
301 |
by (erule PRIMREC.induct) (blast intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)+ |
|
11024 | 302 |
|
27626 | 303 |
theorem ack_not_PRIMREC: |
69880 | 304 |
"\<not> PRIMREC (\<lambda>l. case l of [] \<Rightarrow> 0 | x # l' \<Rightarrow> ack x x)" |
305 |
proof |
|
306 |
assume *: "PRIMREC (\<lambda>l. case l of [] \<Rightarrow> 0 | x # l' \<Rightarrow> ack x x)" |
|
307 |
then obtain m where m: "\<And>l. (case l of [] \<Rightarrow> 0 | x # l' \<Rightarrow> ack x x) < ack m (sum_list l)" |
|
308 |
using ack_bounds_PRIMREC by metis |
|
309 |
show False |
|
310 |
using m [of "[m]"] by simp |
|
311 |
qed |
|
3335 | 312 |
|
313 |
end |