src/ZF/ex/Term.thy
changeset 515 abcc438e7c27
child 740 281881c08397
equal deleted inserted replaced
514:ab2c867829ec 515:abcc438e7c27
       
     1 (*  Title: 	ZF/ex/Term.thy
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1994  University of Cambridge
       
     5 
       
     6 Terms over an alphabet.
       
     7 Illustrates the list functor (essentially the same type as in Trees & Forests)
       
     8 *)
       
     9 
       
    10 Term = List +
       
    11 consts
       
    12   term_rec  :: "[i, [i,i,i]=>i] => i"
       
    13   term_map  :: "[i=>i, i] => i"
       
    14   term_size :: "i=>i"
       
    15   reflect   :: "i=>i"
       
    16   preorder  :: "i=>i"
       
    17 
       
    18   term      :: "i=>i"
       
    19 
       
    20 datatype
       
    21   "term(A)" = Apply ("a: A", "l: list(term(A))")
       
    22   monos	      "[list_mono]"
       
    23   type_intrs  "[list_univ RS subsetD]"
       
    24 
       
    25 rules
       
    26   term_rec_def
       
    27    "term_rec(t,d) == \
       
    28 \   Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z.g`z, zs)), t))"
       
    29 
       
    30   term_map_def	"term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))"
       
    31 
       
    32   term_size_def	"term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))"
       
    33 
       
    34   reflect_def	"reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))"
       
    35 
       
    36   preorder_def	"preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))"
       
    37 
       
    38 end