author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 0 | 7949f97df77a |
child 72 | 30e80f028c57 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/ex/simult |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
Primitives for simultaneous recursive type definitions |
|
7 |
includes worked example of trees & forests |
|
8 |
||
9 |
This is essentially the same data structure that on ex/term.ML, which is |
|
10 |
simpler because it uses List as a new type former. The approach in this |
|
11 |
file may be superior for other simultaneous recursions. |
|
12 |
*) |
|
13 |
||
14 |
Simult = List + |
|
15 |
types tree,forest 1 |
|
16 |
arities tree,forest :: (term)term |
|
17 |
consts |
|
18 |
Part :: "['a set, 'a=>'a] => 'a set" |
|
19 |
TF :: "'a node set set => 'a node set set" |
|
20 |
FNIL :: "'a node set" |
|
21 |
TCONS,FCONS :: "['a node set, 'a node set] => 'a node set" |
|
22 |
Rep_Tree :: "'a tree => 'a node set" |
|
23 |
Abs_Tree :: "'a node set => 'a tree" |
|
24 |
Rep_Forest :: "'a forest => 'a node set" |
|
25 |
Abs_Forest :: "'a node set => 'a forest" |
|
26 |
Tcons :: "['a, 'a forest] => 'a tree" |
|
27 |
Fcons :: "['a tree, 'a forest] => 'a forest" |
|
28 |
Fnil :: "'a forest" |
|
29 |
TF_rec :: "['a node set, ['a node set , 'a node set, 'b]=>'b, \ |
|
30 |
\ 'b, ['a node set , 'a node set, 'b, 'b]=>'b] => 'b" |
|
31 |
tree_rec :: "['a tree, ['a, 'a forest, 'b]=>'b, \ |
|
32 |
\ 'b, ['a tree, 'a forest, 'b, 'b]=>'b] => 'b" |
|
33 |
forest_rec :: "['a forest, ['a, 'a forest, 'b]=>'b, \ |
|
34 |
\ 'b, ['a tree, 'a forest, 'b, 'b]=>'b] => 'b" |
|
35 |
||
36 |
rules |
|
37 |
(*operator for selecting out the various types*) |
|
38 |
Part_def "Part(A,h) == {x. x:A & (? z. x = h(z))}" |
|
39 |
||
40 |
TF_def "TF(A) == lfp(%Z. A <*> Part(Z,In1) \ |
|
41 |
\ <+> ({Numb(0)} <+> Part(Z,In0) <*> Part(Z,In1)))" |
|
42 |
(*faking a type definition for tree...*) |
|
43 |
Rep_Tree "Rep_Tree(n): Part(TF(range(Leaf)),In0)" |
|
44 |
Rep_Tree_inverse "Abs_Tree(Rep_Tree(t)) = t" |
|
45 |
Abs_Tree_inverse "z: Part(TF(range(Leaf)),In0) ==> Rep_Tree(Abs_Tree(z)) = z" |
|
46 |
(*faking a type definition for forest...*) |
|
47 |
Rep_Forest "Rep_Forest(n): Part(TF(range(Leaf)),In1)" |
|
48 |
Rep_Forest_inverse "Abs_Forest(Rep_Forest(ts)) = ts" |
|
49 |
Abs_Forest_inverse |
|
50 |
"z: Part(TF(range(Leaf)),In1) ==> Rep_Forest(Abs_Forest(z)) = z" |
|
51 |
||
52 |
(*the concrete constants*) |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
53 |
TCONS_def "TCONS(M,N) == In0(M $ N)" |
0 | 54 |
FNIL_def "FNIL == In1(NIL)" |
55 |
FCONS_def "FCONS(M,N) == In1(CONS(M,N))" |
|
56 |
(*the abstract constants*) |
|
57 |
Tcons_def "Tcons(a,ts) == Abs_Tree(TCONS(Leaf(a), Rep_Forest(ts)))" |
|
58 |
Fnil_def "Fnil == Abs_Forest(FNIL)" |
|
59 |
Fcons_def "Fcons(t,ts) == Abs_Forest(FCONS(Rep_Tree(t), Rep_Forest(ts)))" |
|
60 |
||
61 |
(*recursion*) |
|
62 |
TF_rec_def |
|
63 |
"TF_rec(M,b,c,d) == wfrec(trancl(pred_Sexp), M, \ |
|
64 |
\ %Z g. Case(Z, %U. Split(U, %x y. b(x,y,g(y))), \ |
|
65 |
\ %V. List_case(V, c, \ |
|
66 |
\ %x y. d(x,y,g(x),g(y)))))" |
|
67 |
||
68 |
tree_rec_def |
|
69 |
"tree_rec(t,b,c,d) == \ |
|
70 |
\ TF_rec(Rep_Tree(t), %x y r. b(Inv(Leaf,x), Abs_Forest(y), r), \ |
|
71 |
\ c, %x y rt rf. d(Abs_Tree(x), Abs_Forest(y), rt, rf))" |
|
72 |
||
73 |
forest_rec_def |
|
74 |
"forest_rec(tf,b,c,d) == \ |
|
75 |
\ TF_rec(Rep_Forest(tf), %x y r. b(Inv(Leaf,x), Abs_Forest(y), r), \ |
|
76 |
\ c, %x y rt rf. d(Abs_Tree(x), Abs_Forest(y), rt, rf))" |
|
77 |
end |