| author | chaieb | 
| Wed, 27 Feb 2008 14:39:51 +0100 | |
| changeset 26157 | 4d9d0a26c32a | 
| parent 26072 | f65a7fa2da6c | 
| child 26334 | 80ec6cf82d95 | 
| 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: 
19736diff
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: 
19736diff
changeset | 49 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
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: 
19736diff
changeset | 53 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
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: 
19736diff
changeset | 57 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
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: 
19736diff
changeset | 61 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
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)" | |
| 26072 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 141 | proof (induct j) | 
| 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 142 | case 0 show ?case by simp | 
| 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 143 | next | 
| 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 144 | case (Suc j) show ?case | 
| 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 145 | by (auto intro!: ack_le_mono2) | 
| 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 146 | (metis Suc Suc_leI Suc_lessI less_ack2 linorder_not_less) | 
| 
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
 haftmann parents: 
25157diff
changeset | 147 | qed | 
| 11024 | 148 | |
| 149 | ||
| 150 | text {* PROPERTY A 7-, the single-step lemma *}
 | |
| 151 | ||
| 152 | lemma ack_less_ack_Suc1 [iff]: "ack (i, j) < ack (Suc i, j)" | |
| 153 | apply (blast intro: ack_less_mono2 less_le_trans) | |
| 154 | done | |
| 155 | ||
| 156 | ||
| 19676 | 157 | text {* PROPERTY A 4'? Extra lemma needed for @{term CONSTANT} case, constant functions *}
 | 
| 11024 | 158 | |
| 159 | lemma less_ack1 [iff]: "i < ack (i, j)" | |
| 160 | apply (induct i) | |
| 161 | apply simp_all | |
| 162 | apply (blast intro: Suc_leI le_less_trans) | |
| 163 | done | |
| 164 | ||
| 165 | ||
| 166 | text {* PROPERTY A 8 *}
 | |
| 167 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 168 | lemma ack_1 [simp]: "ack (Suc 0, j) = j + 2" | 
| 11024 | 169 | apply (induct j) | 
| 170 | apply simp_all | |
| 171 | done | |
| 172 | ||
| 173 | ||
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11464diff
changeset | 174 | text {* PROPERTY A 9.  The unary @{text 1} and @{text 2} in @{term
 | 
| 11024 | 175 | ack} is essential for the rewriting. *} | 
| 176 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 177 | lemma ack_2 [simp]: "ack (Suc (Suc 0), j) = 2 * j + 3" | 
| 11024 | 178 | apply (induct j) | 
| 179 | apply simp_all | |
| 180 | done | |
| 3335 | 181 | |
| 182 | ||
| 11024 | 183 | text {* PROPERTY A 7, monotonicity for @{text "<"} [not clear why
 | 
| 184 |   @{thm [source] ack_1} is now needed first!] *}
 | |
| 185 | ||
| 186 | lemma ack_less_mono1_aux: "ack (i, k) < ack (Suc (i +i'), k)" | |
| 187 | apply (induct i k rule: ack.induct) | |
| 188 | apply simp_all | |
| 189 | prefer 2 | |
| 190 | apply (blast intro: less_trans ack_less_mono2) | |
| 191 | apply (induct_tac i' n rule: ack.induct) | |
| 192 | apply simp_all | |
| 193 | apply (blast intro: Suc_leI [THEN le_less_trans] ack_less_mono2) | |
| 194 | done | |
| 195 | ||
| 196 | lemma ack_less_mono1: "i < j ==> ack (i, k) < ack (j, k)" | |
| 197 | apply (drule less_imp_Suc_add) | |
| 198 | apply (blast intro!: ack_less_mono1_aux) | |
| 199 | done | |
| 200 | ||
| 201 | ||
| 202 | text {* PROPERTY A 7', monotonicity for @{text "\<le>"} *}
 | |
| 203 | ||
| 204 | lemma ack_le_mono1: "i \<le> j ==> ack (i, k) \<le> ack (j, k)" | |
| 205 | apply (simp add: order_le_less) | |
| 206 | apply (blast intro: ack_less_mono1) | |
| 207 | done | |
| 208 | ||
| 209 | ||
| 210 | text {* PROPERTY A 10 *}
 | |
| 211 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 212 | lemma ack_nest_bound: "ack(i1, ack (i2, j)) < ack (2 + (i1 + i2), j)" | 
| 11024 | 213 | apply (simp add: numerals) | 
| 214 | apply (rule ack2_le_ack1 [THEN [2] less_le_trans]) | |
| 215 | apply simp | |
| 216 | apply (rule le_add1 [THEN ack_le_mono1, THEN le_less_trans]) | |
| 217 | apply (rule ack_less_mono1 [THEN ack_less_mono2]) | |
| 218 | apply (simp add: le_imp_less_Suc le_add2) | |
| 219 | done | |
| 220 | ||
| 3335 | 221 | |
| 11024 | 222 | text {* PROPERTY A 11 *}
 | 
| 3335 | 223 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 224 | 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: 
24551diff
changeset | 225 | apply (rule less_trans [of _ "ack (Suc (Suc 0), ack (i1 + i2, j))" _]) | 
| 11024 | 226 | prefer 2 | 
| 227 | apply (rule ack_nest_bound [THEN less_le_trans]) | |
| 228 | apply (simp add: Suc3_eq_add_3) | |
| 229 | apply simp | |
| 230 | apply (cut_tac i = i1 and m1 = i2 and k = j in le_add1 [THEN ack_le_mono1]) | |
| 231 | apply (cut_tac i = "i2" and m1 = i1 and k = j in le_add2 [THEN ack_le_mono1]) | |
| 232 | apply auto | |
| 233 | done | |
| 234 | ||
| 235 | ||
| 236 | text {* PROPERTY A 12.  Article uses existential quantifier but the ALF proof
 | |
| 237 |   used @{text "k + 4"}.  Quantified version must be nested @{text
 | |
| 238 | "\<exists>k'. \<forall>i j. ..."} *} | |
| 3335 | 239 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 240 | 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: 
24551diff
changeset | 241 | 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: 
24551diff
changeset | 242 | apply (blast intro: add_less_mono less_ack2) | 
| 11024 | 243 | apply (rule ack_add_bound [THEN less_le_trans]) | 
| 244 | apply simp | |
| 245 | done | |
| 246 | ||
| 247 | ||
| 248 | ||
| 249 | text {* Inductive definition of the @{term PR} functions *}
 | |
| 3335 | 250 | |
| 11024 | 251 | text {* MAIN RESULT *}
 | 
| 252 | ||
| 253 | lemma SC_case: "SC l < ack (1, list_add l)" | |
| 254 | apply (unfold SC_def) | |
| 255 | apply (induct l) | |
| 256 | apply (simp_all add: le_add1 le_imp_less_Suc) | |
| 257 | done | |
| 258 | ||
| 19676 | 259 | lemma CONSTANT_case: "CONSTANT k l < ack (k, list_add l)" | 
| 24551 | 260 | by simp | 
| 3335 | 261 | |
| 11024 | 262 | lemma PROJ_case [rule_format]: "\<forall>i. PROJ i l < ack (0, list_add l)" | 
| 263 | apply (simp add: PROJ_def) | |
| 264 | apply (induct l) | |
| 24551 | 265 | apply (auto simp add: drop_Cons split: nat.split) | 
| 266 | apply (blast intro: less_le_trans le_add2) | |
| 11024 | 267 | done | 
| 268 | ||
| 269 | ||
| 270 | text {* @{term COMP} case *}
 | |
| 3335 | 271 | |
| 22944 | 272 | lemma COMP_map_aux: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack (kf, list_add l)) | 
| 11024 | 273 | ==> \<exists>k. \<forall>l. list_add (map (\<lambda>f. f l) fs) < ack (k, list_add l)" | 
| 22944 | 274 | apply (induct fs) | 
| 24551 | 275 | apply (rule_tac x = 0 in exI) | 
| 11024 | 276 | apply simp | 
| 277 | apply simp | |
| 278 | apply (blast intro: add_less_mono ack_add_bound less_trans) | |
| 279 | done | |
| 280 | ||
| 281 | lemma COMP_case: | |
| 282 | "\<forall>l. g l < ack (kg, list_add l) ==> | |
| 22944 | 283 | \<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack(kf, list_add l)) | 
| 11024 | 284 | ==> \<exists>k. \<forall>l. COMP g fs l < ack(k, list_add l)" | 
| 285 | apply (unfold COMP_def) | |
| 16588 | 286 |     --{*Now, if meson tolerated map, we could finish with
 | 
| 16731 | 287 |   @{text "(drule COMP_map_aux, meson ack_less_mono2 ack_nest_bound less_trans)"} *}
 | 
| 16588 | 288 | apply (erule COMP_map_aux [THEN exE]) | 
| 289 | apply (rule exI) | |
| 290 | apply (rule allI) | |
| 291 | apply (drule spec)+ | |
| 292 | apply (erule less_trans) | |
| 293 | apply (blast intro: ack_less_mono2 ack_nest_bound less_trans) | |
| 11024 | 294 | done | 
| 295 | ||
| 296 | ||
| 297 | text {* @{term PREC} case *}
 | |
| 3335 | 298 | |
| 11024 | 299 | lemma PREC_case_aux: | 
| 300 | "\<forall>l. f l + list_add l < ack (kf, list_add l) ==> | |
| 301 | \<forall>l. g l + list_add l < ack (kg, list_add l) ==> | |
| 302 | PREC f g l + list_add l < ack (Suc (kf + kg), list_add l)" | |
| 303 | apply (unfold PREC_def) | |
| 304 | apply (case_tac l) | |
| 305 | apply simp_all | |
| 306 | apply (blast intro: less_trans) | |
| 307 |   apply (erule ssubst) -- {* get rid of the needless assumption *}
 | |
| 308 | apply (induct_tac a) | |
| 309 | apply simp_all | |
| 310 |    txt {* base case *}
 | |
| 311 | apply (blast intro: le_add1 [THEN le_imp_less_Suc, THEN ack_less_mono1] less_trans) | |
| 312 |   txt {* induction step *}
 | |
| 313 | apply (rule Suc_leI [THEN le_less_trans]) | |
| 314 | apply (rule le_refl [THEN add_le_mono, THEN le_less_trans]) | |
| 315 | prefer 2 | |
| 316 | apply (erule spec) | |
| 317 | apply (simp add: le_add2) | |
| 318 |   txt {* final part of the simplification *}
 | |
| 319 | apply simp | |
| 320 | apply (rule le_add2 [THEN ack_le_mono1, THEN le_less_trans]) | |
| 321 | apply (erule ack_less_mono2) | |
| 322 | done | |
| 323 | ||
| 324 | lemma PREC_case: | |
| 325 | "\<forall>l. f l < ack (kf, list_add l) ==> | |
| 326 | \<forall>l. g l < ack (kg, list_add l) ==> | |
| 327 | \<exists>k. \<forall>l. PREC f g l < ack (k, list_add l)" | |
| 24551 | 328 | by (metis le_less_trans [OF le_add1 PREC_case_aux] ack_add_bound2) | 
| 11024 | 329 | |
| 22283 | 330 | lemma ack_bounds_PRIMREC: "PRIMREC f ==> \<exists>k. \<forall>l. f l < ack (k, list_add l)" | 
| 11024 | 331 | apply (erule PRIMREC.induct) | 
| 19676 | 332 | apply (blast intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)+ | 
| 11024 | 333 | done | 
| 334 | ||
| 22283 | 335 | lemma ack_not_PRIMREC: "\<not> PRIMREC (\<lambda>l. case l of [] => 0 | x # l' => ack (x, x))" | 
| 11024 | 336 | apply (rule notI) | 
| 337 | 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: 
24551diff
changeset | 338 | apply (rule Nat.less_irrefl) | 
| 11024 | 339 | apply (drule_tac x = "[x]" in spec) | 
| 340 | apply simp | |
| 341 | done | |
| 3335 | 342 | |
| 343 | end |