author | wenzelm |
Fri, 19 Dec 1997 12:09:58 +0100 | |
changeset 4456 | 44e57a6d947d |
parent 3947 | eb707467f8c5 |
child 5191 | 8ceaa19f7717 |
permissions | -rw-r--r-- |
923 | 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 |
Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat) |
|
7 |
||
8 |
Defines "Cartesian Product" and "Disjoint Sum" as set operations. |
|
9 |
Could <*> be generalized to a general summation (Sigma)? |
|
10 |
*) |
|
11 |
||
12 |
Univ = Arith + Sum + |
|
13 |
||
14 |
(** lists, trees will be sets of nodes **) |
|
15 |
||
3947 | 16 |
global |
17 |
||
1475 | 18 |
typedef (Node) |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
19 |
'a node = "{p. EX f x k. p = (f::nat=>nat, x::'a+nat) & f(k)=0}" |
923 | 20 |
|
21 |
types |
|
1384 | 22 |
'a item = 'a node set |
923 | 23 |
|
24 |
consts |
|
25 |
apfst :: "['a=>'c, 'a*'b] => 'c*'b" |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
26 |
Push :: [nat, nat=>nat] => (nat=>nat) |
923 | 27 |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
28 |
Push_Node :: [nat, 'a node] => 'a node |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
29 |
ndepth :: 'a node => nat |
923 | 30 |
|
31 |
Atom :: "('a+nat) => 'a item" |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
32 |
Leaf :: 'a => 'a item |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
33 |
Numb :: nat => 'a item |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
34 |
"$" :: ['a item, 'a item]=> 'a item (infixr 60) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
35 |
In0,In1 :: 'a item => 'a item |
923 | 36 |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
37 |
ntrunc :: [nat, 'a item] => 'a item |
923 | 38 |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
39 |
"<*>" :: ['a item set, 'a item set]=> 'a item set (infixr 80) |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
40 |
"<+>" :: ['a item set, 'a item set]=> 'a item set (infixr 70) |
923 | 41 |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
42 |
Split :: [['a item, 'a item]=>'b, 'a item] => 'b |
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
43 |
Case :: [['a item]=>'b, ['a item]=>'b, 'a item] => 'b |
923 | 44 |
|
45 |
diag :: "'a set => ('a * 'a)set" |
|
1151 | 46 |
"<**>" :: "[('a item * 'a item)set, ('a item * 'a item)set] |
47 |
=> ('a item * 'a item)set" (infixr 80) |
|
48 |
"<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] |
|
49 |
=> ('a item * 'a item)set" (infixr 70) |
|
923 | 50 |
|
3947 | 51 |
|
52 |
local |
|
53 |
||
923 | 54 |
defs |
55 |
||
56 |
Push_Node_def "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))" |
|
57 |
||
58 |
(*crude "lists" of nats -- needed for the constructions*) |
|
1396 | 59 |
apfst_def "apfst == (%f (x,y). (f(x),y))" |
923 | 60 |
Push_def "Push == (%b h. nat_case (Suc b) h)" |
61 |
||
62 |
(** operations on S-expressions -- sets of nodes **) |
|
63 |
||
64 |
(*S-expression constructors*) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
65 |
Atom_def "Atom == (%x. {Abs_Node((%k.0, x))})" |
923 | 66 |
Scons_def "M$N == (Push_Node(0) `` M) Un (Push_Node(Suc(0)) `` N)" |
67 |
||
68 |
(*Leaf nodes, with arbitrary or nat labels*) |
|
69 |
Leaf_def "Leaf == Atom o Inl" |
|
70 |
Numb_def "Numb == Atom o Inr" |
|
71 |
||
72 |
(*Injections of the "disjoint sum"*) |
|
73 |
In0_def "In0(M) == Numb(0) $ M" |
|
74 |
In1_def "In1(M) == Numb(Suc(0)) $ M" |
|
75 |
||
76 |
(*the set of nodes with depth less than k*) |
|
1068 | 77 |
ndepth_def "ndepth(n) == (%(f,x). LEAST k. f(k)=0) (Rep_Node n)" |
923 | 78 |
ntrunc_def "ntrunc k N == {n. n:N & ndepth(n)<k}" |
79 |
||
80 |
(*products and sums for the "universe"*) |
|
81 |
uprod_def "A<*>B == UN x:A. UN y:B. { (x$y) }" |
|
82 |
usum_def "A<+>B == In0``A Un In1``B" |
|
83 |
||
84 |
(*the corresponding eliminators*) |
|
85 |
Split_def "Split c M == @u. ? x y. M = x$y & u = c x y" |
|
86 |
||
1151 | 87 |
Case_def "Case c d M == @u. (? x . M = In0(x) & u = c(x)) |
88 |
| (? y . M = In1(y) & u = d(y))" |
|
923 | 89 |
|
90 |
||
91 |
(** diagonal sets and equality for the "universe" **) |
|
92 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
923
diff
changeset
|
93 |
diag_def "diag(A) == UN x:A. {(x,x)}" |
923 | 94 |
|
1068 | 95 |
dprod_def "r<**>s == UN (x,x'):r. UN (y,y'):s. {(x$y,x'$y')}" |
923 | 96 |
|
1151 | 97 |
dsum_def "r<++>s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un |
98 |
(UN (y,y'):s. {(In1(y),In1(y'))})" |
|
923 | 99 |
|
100 |
end |