author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 0 | 7949f97df77a |
child 51 | 934a58983311 |
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 |
||
14 |
Univ = Arith + |
|
15 |
types node 1 |
|
16 |
arities node :: (term)term |
|
17 |
consts |
|
18 |
Least :: "(nat=>bool) => nat" (binder "LEAST " 10) |
|
19 |
||
20 |
apfst :: "['a=>'c, 'a*'b] => 'c*'b" |
|
21 |
Push :: "[nat, nat=>nat] => (nat=>nat)" |
|
22 |
||
23 |
Node :: "((nat=>nat) * ('a+nat)) set" |
|
24 |
Rep_Node :: "'a node => (nat=>nat) * ('a+nat)" |
|
25 |
Abs_Node :: "(nat=>nat) * ('a+nat) => 'a node" |
|
26 |
Push_Node :: "[nat, 'a node] => 'a node" |
|
27 |
ndepth :: "'a node => nat" |
|
28 |
||
29 |
Atom :: "('a+nat) => 'a node set" |
|
30 |
Leaf :: "'a => 'a node set" |
|
31 |
Numb :: "nat => 'a node set" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
32 |
"$" :: "['a node set, 'a node set]=> 'a node set" (infixr 60) |
0 | 33 |
In0,In1 :: "'a node set => 'a node set" |
34 |
||
35 |
ntrunc :: "[nat, 'a node set] => 'a node set" |
|
36 |
||
37 |
"<*>" :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 80) |
|
38 |
"<+>" :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 70) |
|
39 |
||
40 |
Split :: "['a node set, ['a node set, 'a node set]=>'b] => 'b" |
|
41 |
Case :: "['a node set, ['a node set]=>'b, ['a node set]=>'b] => 'b" |
|
42 |
||
43 |
diag :: "'a set => ('a * 'a)set" |
|
44 |
"<**>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \ |
|
45 |
\ => ('a node set * 'a node set)set" (infixr 80) |
|
46 |
"<++>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \ |
|
47 |
\ => ('a node set * 'a node set)set" (infixr 70) |
|
48 |
||
49 |
rules |
|
50 |
||
51 |
(*least number operator*) |
|
52 |
Least_def "Least(P) == @k. P(k) & (ALL j. j<k --> ~P(j))" |
|
53 |
||
54 |
(** lists, trees will be sets of nodes **) |
|
55 |
Node_def "Node == {p. EX f x k. p = <f,x> & f(k)=0}" |
|
56 |
(*faking the type definition 'a node == (nat=>nat) * ('a+nat) *) |
|
57 |
Rep_Node "Rep_Node(n): Node" |
|
58 |
Rep_Node_inverse "Abs_Node(Rep_Node(n)) = n" |
|
59 |
Abs_Node_inverse "p: Node ==> Rep_Node(Abs_Node(p)) = p" |
|
60 |
Push_Node_def "Push_Node == (%n x. Abs_Node (apfst(Push(n),Rep_Node(x))))" |
|
61 |
||
62 |
(*crude "lists" of nats -- needed for the constructions*) |
|
63 |
apfst_def "apfst == (%f p.split(p, %x y. <f(x),y>))" |
|
64 |
Push_def "Push == (%b h n. nat_case(n,Suc(b),h))" |
|
65 |
||
66 |
(** operations on S-expressions -- sets of nodes **) |
|
67 |
||
68 |
(*S-expression constructors*) |
|
69 |
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
|
70 |
Scons_def "M$N == (Push_Node(0) `` M) Un (Push_Node(Suc(0)) `` N)" |
0 | 71 |
|
72 |
(*Leaf nodes, with arbitrary or nat labels*) |
|
73 |
Leaf_def "Leaf == Atom o Inl" |
|
74 |
Numb_def "Numb == Atom o Inr" |
|
75 |
||
76 |
(*Injections of the "disjoint sum"*) |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
77 |
In0_def "In0(M) == Numb(0) $ M" |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
78 |
In1_def "In1(M) == Numb(Suc(0)) $ M" |
0 | 79 |
|
80 |
(*the set of nodes with depth less than k*) |
|
81 |
ndepth_def "ndepth(n) == split(Rep_Node(n), %f x. LEAST k. f(k)=0)" |
|
82 |
ntrunc_def "ntrunc(k,N) == {n. n:N & ndepth(n)<k}" |
|
83 |
||
84 |
(*products and sums for the "universe"*) |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
85 |
uprod_def "A<*>B == UN x:A. UN y:B. { (x$y) }" |
0 | 86 |
usum_def "A<+>B == In0``A Un In1``B" |
87 |
||
88 |
(*the corresponding eliminators*) |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
89 |
Split_def "Split(M,c) == @u. ? x y. M = x$y & u = c(x,y)" |
0 | 90 |
|
91 |
Case_def "Case(M,c,d) == @u. (? x . M = In0(x) & u = c(x)) \ |
|
92 |
\ | (? y . M = In1(y) & u = d(y))" |
|
93 |
||
94 |
||
95 |
(** diagonal sets and equality for the "universe" **) |
|
96 |
||
97 |
diag_def "diag(A) == UN x:A. {<x,x>}" |
|
98 |
||
99 |
dprod_def "r<**>s == UN u:r. UN v:s. \ |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
100 |
\ split(u, %x x'. split(v, %y y'. {<x$y,x'$y'>}))" |
0 | 101 |
|
102 |
dsum_def "r<++>s == (UN u:r. split(u, %x x'. {<In0(x),In0(x')>})) Un \ |
|
103 |
\ (UN v:s. split(v, %y y'. {<In1(y),In1(y')>}))" |
|
104 |
||
105 |
end |