author | clasohm |
Wed, 02 Nov 1994 11:50:09 +0100 | |
changeset 156 | fd1be45b64bf |
parent 128 | 89669c58e506 |
child 190 | 5505c746fff7 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/univ.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
Move LEAST to nat.thy??? Could it be defined for all types 'a::ord? |
|
7 |
||
8 |
Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat) |
|
9 |
||
10 |
Defines "Cartesian Product" and "Disjoint Sum" as set operations. |
|
11 |
Could <*> be generalized to a general summation (Sigma)? |
|
12 |
*) |
|
13 |
||
111 | 14 |
Univ = Arith + Sum + |
51 | 15 |
|
16 |
types |
|
17 |
'a node |
|
128 | 18 |
'a item = "'a node set" |
51 | 19 |
|
20 |
arities |
|
21 |
node :: (term)term |
|
22 |
||
0 | 23 |
consts |
24 |
Least :: "(nat=>bool) => nat" (binder "LEAST " 10) |
|
25 |
||
26 |
apfst :: "['a=>'c, 'a*'b] => 'c*'b" |
|
27 |
Push :: "[nat, nat=>nat] => (nat=>nat)" |
|
28 |
||
29 |
Node :: "((nat=>nat) * ('a+nat)) set" |
|
30 |
Rep_Node :: "'a node => (nat=>nat) * ('a+nat)" |
|
31 |
Abs_Node :: "(nat=>nat) * ('a+nat) => 'a node" |
|
32 |
Push_Node :: "[nat, 'a node] => 'a node" |
|
33 |
ndepth :: "'a node => nat" |
|
34 |
||
128 | 35 |
Atom :: "('a+nat) => 'a item" |
36 |
Leaf :: "'a => 'a item" |
|
37 |
Numb :: "nat => 'a item" |
|
38 |
"$" :: "['a item, 'a item]=> 'a item" (infixr 60) |
|
39 |
In0,In1 :: "'a item => 'a item" |
|
0 | 40 |
|
128 | 41 |
ntrunc :: "[nat, 'a item] => 'a item" |
0 | 42 |
|
128 | 43 |
"<*>" :: "['a item set, 'a item set]=> 'a item set" (infixr 80) |
44 |
"<+>" :: "['a item set, 'a item set]=> 'a item set" (infixr 70) |
|
0 | 45 |
|
128 | 46 |
Split :: "[['a item, 'a item]=>'b, 'a item] => 'b" |
47 |
Case :: "[['a item]=>'b, ['a item]=>'b, 'a item] => 'b" |
|
0 | 48 |
|
49 |
diag :: "'a set => ('a * 'a)set" |
|
128 | 50 |
"<**>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \ |
51 |
\ => ('a item * 'a item)set" (infixr 80) |
|
52 |
"<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \ |
|
53 |
\ => ('a item * 'a item)set" (infixr 70) |
|
0 | 54 |
|
55 |
rules |
|
56 |
||
57 |
(*least number operator*) |
|
58 |
Least_def "Least(P) == @k. P(k) & (ALL j. j<k --> ~P(j))" |
|
59 |
||
60 |
(** lists, trees will be sets of nodes **) |
|
61 |
Node_def "Node == {p. EX f x k. p = <f,x> & f(k)=0}" |
|
62 |
(*faking the type definition 'a node == (nat=>nat) * ('a+nat) *) |
|
63 |
Rep_Node "Rep_Node(n): Node" |
|
64 |
Rep_Node_inverse "Abs_Node(Rep_Node(n)) = n" |
|
65 |
Abs_Node_inverse "p: Node ==> Rep_Node(Abs_Node(p)) = p" |
|
66 |
Push_Node_def "Push_Node == (%n x. Abs_Node (apfst(Push(n),Rep_Node(x))))" |
|
67 |
||
68 |
(*crude "lists" of nats -- needed for the constructions*) |
|
111 | 69 |
apfst_def "apfst == (%f. split(%x y. <f(x),y>))" |
70 |
Push_def "Push == (%b h. nat_case(Suc(b),h))" |
|
0 | 71 |
|
72 |
(** operations on S-expressions -- sets of nodes **) |
|
73 |
||
74 |
(*S-expression constructors*) |
|
75 |
Atom_def "Atom == (%x. {Abs_Node(<%k.0, x>)})" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
76 |
Scons_def "M$N == (Push_Node(0) `` M) Un (Push_Node(Suc(0)) `` N)" |
0 | 77 |
|
78 |
(*Leaf nodes, with arbitrary or nat labels*) |
|
79 |
Leaf_def "Leaf == Atom o Inl" |
|
80 |
Numb_def "Numb == Atom o Inr" |
|
81 |
||
82 |
(*Injections of the "disjoint sum"*) |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
83 |
In0_def "In0(M) == Numb(0) $ M" |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
84 |
In1_def "In1(M) == Numb(Suc(0)) $ M" |
0 | 85 |
|
86 |
(*the set of nodes with depth less than k*) |
|
111 | 87 |
ndepth_def "ndepth(n) == split(%f x. LEAST k. f(k)=0, Rep_Node(n))" |
0 | 88 |
ntrunc_def "ntrunc(k,N) == {n. n:N & ndepth(n)<k}" |
89 |
||
90 |
(*products and sums for the "universe"*) |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
91 |
uprod_def "A<*>B == UN x:A. UN y:B. { (x$y) }" |
0 | 92 |
usum_def "A<+>B == In0``A Un In1``B" |
93 |
||
94 |
(*the corresponding eliminators*) |
|
111 | 95 |
Split_def "Split(c,M) == @u. ? x y. M = x$y & u = c(x,y)" |
0 | 96 |
|
111 | 97 |
Case_def "Case(c,d,M) == @u. (? x . M = In0(x) & u = c(x)) \ |
0 | 98 |
\ | (? y . M = In1(y) & u = d(y))" |
99 |
||
100 |
||
101 |
(** diagonal sets and equality for the "universe" **) |
|
102 |
||
103 |
diag_def "diag(A) == UN x:A. {<x,x>}" |
|
104 |
||
111 | 105 |
dprod_def "r<**>s == UN u:r. split(%x x'. \ |
106 |
\ UN v:s. split(%y y'. {<x$y,x'$y'>}, v), u)" |
|
0 | 107 |
|
111 | 108 |
dsum_def "r<++>s == (UN u:r. split(%x x'. {<In0(x),In0(x')>}, u)) Un \ |
109 |
\ (UN v:s. split(%y y'. {<In1(y),In1(y')>}, v))" |
|
0 | 110 |
|
111 |
end |