src/ZF/ex/TF_Fn.thy
changeset 13895 b6105462ccd3
parent 13894 8018173a7979
child 13896 717bd79b976f
equal deleted inserted replaced
13894:8018173a7979 13895:b6105462ccd3
     1 (*  Title: 	ZF/ex/TF.thy
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1993  University of Cambridge
       
     5 
       
     6 Trees & forests, a mutually recursive type definition.
       
     7 *)
       
     8 
       
     9 TF_Fn = TF + ListFn +
       
    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 rules
       
    19   TF_rec_def
       
    20     "TF_rec(z,b,c,d) == Vrec(z,  			\
       
    21 \      %z r. tree_forest_case(%x f. b(x, f, r`f), 	\
       
    22 \                             c, 			\
       
    23 \		              %t f. d(t, f, r`t, r`f), z))"
       
    24 
       
    25   list_of_TF_def
       
    26     "list_of_TF(z) == TF_rec(z, %x f r. [Tcons(x,f)], [], \
       
    27 \		             %t f r1 r2. Cons(t, r2))"
       
    28 
       
    29   TF_of_list_def
       
    30     "TF_of_list(f) == list_rec(f, Fnil,  %t f r. Fcons(t,r))"
       
    31 
       
    32   TF_map_def
       
    33     "TF_map(h,z) == TF_rec(z, %x f r.Tcons(h(x),r), Fnil, \
       
    34 \                           %t f r1 r2. Fcons(r1,r2))"
       
    35 
       
    36   TF_size_def
       
    37     "TF_size(z) == TF_rec(z, %x f r.succ(r), 0, %t f r1 r2. r1#+r2)"
       
    38 
       
    39   TF_preorder_def
       
    40     "TF_preorder(z) == TF_rec(z, %x f r.Cons(x,r), Nil, %t f r1 r2. r1@r2)"
       
    41 
       
    42 end