| author | wenzelm |
| Tue, 16 Oct 2001 17:56:12 +0200 | |
| changeset 11808 | c724a9093ebe |
| parent 11354 | 9b80fe19407f |
| permissions | -rw-r--r-- |
| 1478 | 1 |
(* Title: ZF/ex/Term.thy |
| 515 | 2 |
ID: $Id$ |
| 1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
| 515 | 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 |
||
|
11354
9b80fe19407f
examples files start from Main instead of various ZF theories
paulson
parents:
11316
diff
changeset
|
10 |
Term = Main + |
| 515 | 11 |
consts |
| 1401 | 12 |
term :: i=>i |
| 515 | 13 |
|
14 |
datatype |
|
| 11316 | 15 |
"term(A)" = Apply ("a \\<in> A", "l \\<in> list(term(A))")
|
| 6117 | 16 |
monos list_mono |
| 740 | 17 |
|
18 |
type_elims "[make_elim (list_univ RS subsetD)]" |
|
19 |
(*Or could have |
|
20 |
type_intrs "[list_univ RS subsetD]" |
|
21 |
*) |
|
| 515 | 22 |
|
| 6159 | 23 |
constdefs |
| 6112 | 24 |
term_rec :: [i, [i,i,i]=>i] => i |
| 1155 | 25 |
"term_rec(t,d) == |
| 6159 | 26 |
Vrec(t, %t g. term_case(%x zs. d(x, zs, map(%z. g`z, zs)), t))" |
| 515 | 27 |
|
| 6159 | 28 |
term_map :: [i=>i, i] => i |
29 |
"term_map(f,t) == term_rec(t, %x zs rs. Apply(f(x), rs))" |
|
| 515 | 30 |
|
| 6159 | 31 |
term_size :: i=>i |
32 |
"term_size(t) == term_rec(t, %x zs rs. succ(list_add(rs)))" |
|
| 515 | 33 |
|
| 6159 | 34 |
reflect :: i=>i |
35 |
"reflect(t) == term_rec(t, %x zs rs. Apply(x, rev(rs)))" |
|
| 515 | 36 |
|
| 6159 | 37 |
preorder :: i=>i |
38 |
"preorder(t) == term_rec(t, %x zs rs. Cons(x, flat(rs)))" |
|
| 515 | 39 |
|
| 6159 | 40 |
postorder :: i=>i |
41 |
"postorder(t) == term_rec(t, %x zs rs. flat(rs) @ [x])" |
|
| 6112 | 42 |
|
| 515 | 43 |
end |