src/ZF/Induct/Primrec_defs.thy
author wenzelm
Wed, 05 Dec 2001 15:45:24 +0100
changeset 12393 03c55bb0ee92
parent 12088 6f463d16cbd0
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ex/Primrec_defs.thy
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     2
    ID:         $Id$
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     5
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     6
Primitive Recursive Functions: preliminary definitions of the constructors
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     7
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     8
[These must reside in a separate theory in order to be visible in the
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     9
 con_defs part of prim_rec's inductive definition.]
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    10
*)
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    11
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    12
Primrec_defs = Main +
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    13
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    14
consts
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    15
    SC      :: i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    16
    CONST   :: i=>i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    17
    PROJ    :: i=>i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    18
    COMP    :: [i,i]=>i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    19
    PREC    :: [i,i]=>i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    20
    ACK     :: i=>i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    21
    ack     :: [i,i]=>i
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    22
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    23
translations
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    24
  "ack(x,y)"  == "ACK(x) ` [y]"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    25
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    26
defs
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    27
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    28
  SC_def    "SC == \\<lambda>l \\<in> list(nat).list_case(0, %x xs. succ(x), l)"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    29
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    30
  CONST_def "CONST(k) == \\<lambda>l \\<in> list(nat).k"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    31
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    32
  PROJ_def  "PROJ(i) == \\<lambda>l \\<in> list(nat). list_case(0, %x xs. x, drop(i,l))"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    33
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    34
  COMP_def  "COMP(g,fs) == \\<lambda>l \\<in> list(nat). g ` List.map(%f. f`l, fs)"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    35
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    36
  (*Note that g is applied first to PREC(f,g)`y and then to y!*)
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    37
  PREC_def  "PREC(f,g) == 
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    38
            \\<lambda>l \\<in> list(nat). list_case(0, 
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    39
                      %x xs. rec(x, f`xs, %y r. g ` Cons(r, Cons(y, xs))), l)"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    40
  
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    41
primrec
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    42
  ACK_0     "ACK(0) = SC"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    43
  ACK_succ  "ACK(succ(i)) = PREC (CONST (ACK(i) ` [1]),
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    44
				  COMP(ACK(i), [PROJ(0)]))"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    45
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    46
end