src/ZF/ex/TF.thy
changeset 515 abcc438e7c27
child 581 465075fd257b
equal deleted inserted replaced
514:ab2c867829ec 515:abcc438e7c27
       
     1 (*  Title: 	ZF/ex/TF.thy
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1994  University of Cambridge
       
     5 
       
     6 Trees & forests, a mutually recursive type definition.
       
     7 *)
       
     8 
       
     9 TF = List +
       
    10 consts
       
    11   TF_rec      :: "[i, [i,i,i]=>i, i, [i,i,i,i]=>i] => i"
       
    12   TF_map      :: "[i=>i, i] => i"
       
    13   TF_size     :: "i=>i"
       
    14   TF_preorder :: "i=>i"
       
    15   list_of_TF  :: "i=>i"
       
    16   TF_of_list  :: "i=>i"
       
    17 
       
    18   tree, forest, tree_forest    :: "i=>i"
       
    19 
       
    20 datatype
       
    21   "tree(A)"   = "Tcons" ("a: A",  "f: forest(A)")
       
    22 and
       
    23   "forest(A)" = "Fnil"  |  "Fcons" ("t: tree(A)",  "f: forest(A)")
       
    24 
       
    25 rules
       
    26   TF_rec_def
       
    27     "TF_rec(z,b,c,d) == Vrec(z,  			\
       
    28 \      %z r. tree_forest_case(%x f. b(x, f, r`f), 	\
       
    29 \                             c, 			\
       
    30 \		              %t f. d(t, f, r`t, r`f), z))"
       
    31 
       
    32   list_of_TF_def
       
    33     "list_of_TF(z) == TF_rec(z, %x f r. [Tcons(x,f)], [], \
       
    34 \		             %t f r1 r2. Cons(t, r2))"
       
    35 
       
    36   TF_of_list_def
       
    37     "TF_of_list(f) == list_rec(f, Fnil,  %t f r. Fcons(t,r))"
       
    38 
       
    39   TF_map_def
       
    40     "TF_map(h,z) == TF_rec(z, %x f r.Tcons(h(x),r), Fnil, \
       
    41 \                           %t f r1 r2. Fcons(r1,r2))"
       
    42 
       
    43   TF_size_def
       
    44     "TF_size(z) == TF_rec(z, %x f r.succ(r), 0, %t f r1 r2. r1#+r2)"
       
    45 
       
    46   TF_preorder_def
       
    47     "TF_preorder(z) == TF_rec(z, %x f r.Cons(x,r), Nil, %t f r1 r2. r1@r2)"
       
    48 
       
    49 end