src/ZF/ex/TermFn.thy
changeset 0 a5a9c433f639
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/ex/TermFn.thy	Thu Sep 16 12:20:38 1993 +0200
@@ -0,0 +1,31 @@
+(*  Title: 	ZF/ex/term-fn.thy
+    ID:         $Id$
+    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1993  University of Cambridge
+
+Terms over an alphabet.
+Illustrates the list functor (essentially the same type as in Trees & Forests)
+*)
+
+TermFn = Term + ListFn +
+consts
+    term_rec    :: "[i, [i,i,i]=>i] => i"
+    term_map    :: "[i=>i, i] => i"
+    term_size   :: "i=>i"
+    reflect     :: "i=>i"
+    preorder    :: "i=>i"
+
+rules
+  term_rec_def
+   "term_rec(t,d) == \
+\   Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z.g`z, zs)), t))"
+
+  term_map_def	"term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))"
+
+  term_size_def	"term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))"
+
+  reflect_def	"reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))"
+
+  preorder_def	"preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))"
+
+end