| author | nipkow | 
| Mon, 28 Aug 2000 09:32:51 +0200 | |
| changeset 9689 | 751fde5307e4 | 
| parent 6117 | f9aad8ccd590 | 
| child 11316 | b4e71bd751e4 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/ex/Primrec.thy | 
| 515 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 515 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 6044 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 6 | Primitive Recursive Functions: the inductive definition | 
| 515 | 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 | ||
| 6044 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 17 | Primrec = Primrec_defs + | 
| 515 | 18 | consts | 
| 6044 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 19 | prim_rec :: i | 
| 515 | 20 | |
| 21 | inductive | |
| 6044 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 22 | domains "prim_rec" <= "list(nat)->nat" | 
| 515 | 23 | intrs | 
| 6044 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 24 | SC "SC : prim_rec" | 
| 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 25 | CONST "k: nat ==> CONST(k) : prim_rec" | 
| 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 26 | PROJ "i: nat ==> PROJ(i) : prim_rec" | 
| 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 27 | COMP "[| g: prim_rec; fs: list(prim_rec) |] ==> COMP(g,fs): prim_rec" | 
| 
e0f9d930e956
Needs separate theory Primrec_defs due to new inductive defs package
 paulson parents: 
3841diff
changeset | 28 | PREC "[| f: prim_rec; g: prim_rec |] ==> PREC(f,g): prim_rec" | 
| 6117 | 29 | monos list_mono | 
| 30 | con_defs SC_def, CONST_def, PROJ_def, COMP_def, PREC_def | |
| 3841 | 31 | type_intrs "nat_typechecks @ list.intrs @ | 
| 1478 | 32 | [lam_type, list_case_type, drop_type, map_type, | 
| 33 | apply_type, rec_type]" | |
| 515 | 34 | |
| 35 | end |