author | paulson |
Thu, 31 May 2001 18:28:23 +0200 | |
changeset 11354 | 9b80fe19407f |
parent 11316 | b4e71bd751e4 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/ex/TF.thy |
515 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
515 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Trees & forests, a mutually recursive type definition. |
|
7 |
*) |
|
8 |
||
11354
9b80fe19407f
examples files start from Main instead of various ZF theories
paulson
parents:
11316
diff
changeset
|
9 |
TF = Main + |
515 | 10 |
consts |
6112 | 11 |
tree, forest, tree_forest :: i=>i |
12 |
||
13 |
datatype |
|
11316 | 14 |
"tree(A)" = Tcons ("a \\<in> A", "f \\<in> forest(A)") |
6112 | 15 |
and |
11316 | 16 |
"forest(A)" = Fnil | Fcons ("t \\<in> tree(A)", "f \\<in> forest(A)") |
6112 | 17 |
|
18 |
||
19 |
consts |
|
6046 | 20 |
map :: [i=>i, i] => i |
21 |
size :: i=>i |
|
22 |
preorder :: i=>i |
|
1401 | 23 |
list_of_TF :: i=>i |
6046 | 24 |
of_list :: i=>i |
25 |
reflect :: i=>i |
|
515 | 26 |
|
6046 | 27 |
primrec |
28 |
"list_of_TF (Tcons(x,f)) = [Tcons(x,f)]" |
|
29 |
"list_of_TF (Fnil) = []" |
|
30 |
"list_of_TF (Fcons(t,tf)) = Cons (t, list_of_TF(tf))" |
|
515 | 31 |
|
6046 | 32 |
primrec |
33 |
"of_list([]) = Fnil" |
|
34 |
"of_list(Cons(t,l)) = Fcons(t, of_list(l))" |
|
35 |
||
36 |
primrec |
|
37 |
"map (h, Tcons(x,f)) = Tcons(h(x), map(h,f))" |
|
38 |
"map (h, Fnil) = Fnil" |
|
39 |
"map (h, Fcons(t,tf)) = Fcons (map(h, t), map(h, tf))" |
|
515 | 40 |
|
6046 | 41 |
primrec |
42 |
"size (Tcons(x,f)) = succ(size(f))" |
|
43 |
"size (Fnil) = 0" |
|
44 |
"size (Fcons(t,tf)) = size(t) #+ size(tf)" |
|
45 |
||
46 |
primrec |
|
47 |
"preorder (Tcons(x,f)) = Cons(x, preorder(f))" |
|
48 |
"preorder (Fnil) = Nil" |
|
49 |
"preorder (Fcons(t,tf)) = preorder(t) @ preorder(tf)" |
|
50 |
||
51 |
primrec |
|
52 |
"reflect (Tcons(x,f)) = Tcons(x, reflect(f))" |
|
53 |
"reflect (Fnil) = Fnil" |
|
54 |
"reflect (Fcons(t,tf)) = of_list |
|
55 |
(list_of_TF (reflect(tf)) @ |
|
56 |
Cons(reflect(t), Nil))" |
|
515 | 57 |
|
58 |
end |