src/ZF/ex/Primrec.thy
author wenzelm
Fri, 07 Mar 1997 11:48:46 +0100
changeset 2754 59bd96046ad6
parent 1478 2b8c2a7547ab
child 3840 e0baea4d485a
permissions -rw-r--r--
moved settings comment to build;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/ex/Primrec.thy
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     6
Primitive Recursive Functions
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
Proof adopted from
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
Nora Szasz, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
A Machine Checked Proof that Ackermann's Function is not Primitive Recursive,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
In: Huet & Plotkin, eds., Logical Environments (CUP, 1993), 317-338.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
See also E. Mendelson, Introduction to Mathematical Logic.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
(Van Nostrand, 1964), page 250, exercise 11.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
Primrec = List +
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    19
    primrec :: i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    20
    SC      :: i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    21
    CONST   :: i=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    22
    PROJ    :: i=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    23
    COMP    :: [i,i]=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    24
    PREC    :: [i,i]=>i
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    25
    ACK     :: i=>i
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    26
    ack     :: [i,i]=>i
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
translations
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
  "ack(x,y)"  == "ACK(x) ` [y]"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 515
diff changeset
    31
defs
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
  SC_def    "SC == lam l:list(nat).list_case(0, %x xs.succ(x), l)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    34
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    35
  CONST_def "CONST(k) == lam l:list(nat).k"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    36
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    37
  PROJ_def  "PROJ(i) == lam l:list(nat). list_case(0, %x xs.x, drop(i,l))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
  COMP_def  "COMP(g,fs) == lam l:list(nat). g ` map(%f. f`l, fs)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    40
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
  (*Note that g is applied first to PREC(f,g)`y and then to y!*)
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 753
diff changeset
    42
  PREC_def  "PREC(f,g) == 
928a16e02f9f removed \...\ inside strings
clasohm
parents: 753
diff changeset
    43
            lam l:list(nat). list_case(0, 
928a16e02f9f removed \...\ inside strings
clasohm
parents: 753
diff changeset
    44
                      %x xs. rec(x, f`xs, %y r. g ` Cons(r, Cons(y, xs))), l)"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    45
  
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 753
diff changeset
    46
  ACK_def   "ACK(i) == rec(i, SC, 
928a16e02f9f removed \...\ inside strings
clasohm
parents: 753
diff changeset
    47
                      %z r. PREC (CONST (r`[1]), COMP(r,[PROJ(0)])))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    48
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    49
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    50
inductive
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    51
  domains "primrec" <= "list(nat)->nat"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    52
  intrs
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    53
    SC       "SC : primrec"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    54
    CONST    "k: nat ==> CONST(k) : primrec"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    55
    PROJ     "i: nat ==> PROJ(i) : primrec"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    56
    COMP     "[| g: primrec; fs: list(primrec) |] ==> COMP(g,fs): primrec"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    57
    PREC     "[| f: primrec; g: primrec |] ==> PREC(f,g): primrec"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    58
  monos      "[list_mono]"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    59
  con_defs   "[SC_def,CONST_def,PROJ_def,COMP_def,PREC_def]"
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    60
  type_intrs "nat_typechecks @ list.intrs @                     
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    61
              [lam_type, list_case_type, drop_type, map_type,   
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    62
              apply_type, rec_type]"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    63
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    64
end