515
|
1 |
(* Title: ZF/ex/Primrec.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
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 |
Primrec = List +
|
|
18 |
consts
|
|
19 |
primrec :: "i"
|
|
20 |
SC :: "i"
|
|
21 |
CONST :: "i=>i"
|
|
22 |
PROJ :: "i=>i"
|
|
23 |
COMP :: "[i,i]=>i"
|
|
24 |
PREC :: "[i,i]=>i"
|
|
25 |
ACK :: "i=>i"
|
|
26 |
ack :: "[i,i]=>i"
|
|
27 |
|
|
28 |
translations
|
|
29 |
"ack(x,y)" == "ACK(x) ` [y]"
|
|
30 |
|
753
|
31 |
defs
|
515
|
32 |
|
|
33 |
SC_def "SC == lam l:list(nat).list_case(0, %x xs.succ(x), l)"
|
|
34 |
|
|
35 |
CONST_def "CONST(k) == lam l:list(nat).k"
|
|
36 |
|
|
37 |
PROJ_def "PROJ(i) == lam l:list(nat). list_case(0, %x xs.x, drop(i,l))"
|
|
38 |
|
|
39 |
COMP_def "COMP(g,fs) == lam l:list(nat). g ` map(%f. f`l, fs)"
|
|
40 |
|
|
41 |
(*Note that g is applied first to PREC(f,g)`y and then to y!*)
|
1155
|
42 |
PREC_def "PREC(f,g) ==
|
|
43 |
lam l:list(nat). list_case(0,
|
|
44 |
%x xs. rec(x, f`xs, %y r. g ` Cons(r, Cons(y, xs))), l)"
|
515
|
45 |
|
1155
|
46 |
ACK_def "ACK(i) == rec(i, SC,
|
|
47 |
%z r. PREC (CONST (r`[1]), COMP(r,[PROJ(0)])))"
|
515
|
48 |
|
|
49 |
|
|
50 |
inductive
|
|
51 |
domains "primrec" <= "list(nat)->nat"
|
|
52 |
intrs
|
|
53 |
SC "SC : primrec"
|
|
54 |
CONST "k: nat ==> CONST(k) : primrec"
|
|
55 |
PROJ "i: nat ==> PROJ(i) : primrec"
|
|
56 |
COMP "[| g: primrec; fs: list(primrec) |] ==> COMP(g,fs): primrec"
|
|
57 |
PREC "[| f: primrec; g: primrec |] ==> PREC(f,g): primrec"
|
|
58 |
monos "[list_mono]"
|
|
59 |
con_defs "[SC_def,CONST_def,PROJ_def,COMP_def,PREC_def]"
|
1155
|
60 |
type_intrs "nat_typechecks @ list.intrs @
|
|
61 |
[lam_type, list_case_type, drop_type, map_type,
|
|
62 |
apply_type, rec_type]"
|
515
|
63 |
|
|
64 |
end
|