author | paulson |
Fri, 05 Nov 1999 12:45:37 +0100 | |
changeset 7999 | 7acf6eb8eec1 |
parent 7255 | 853bdbe9973d |
child 8735 | bb2250ac9557 |
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 |
||
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
6 |
Declares the type ('a, 'b) node, a subtype of (nat=>'b+nat) * ('a+nat) |
923 | 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 |
||
7131 | 14 |
setup arith_setup |
15 |
||
16 |
||
923 | 17 |
(** lists, trees will be sets of nodes **) |
18 |
||
3947 | 19 |
global |
20 |
||
1475 | 21 |
typedef (Node) |
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
22 |
('a, 'b) node = "{p. EX f x k. p = (f::nat=>'b+nat, x::'a+nat) & f k = Inr 0}" |
923 | 23 |
|
24 |
types |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
25 |
'a item = ('a, unit) node set |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
26 |
('a, 'b) dtree = ('a, 'b) node set |
923 | 27 |
|
28 |
consts |
|
29 |
apfst :: "['a=>'c, 'a*'b] => 'c*'b" |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
30 |
Push :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))" |
923 | 31 |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
32 |
Push_Node :: "[('b + nat), ('a, 'b) node] => ('a, 'b) node" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
33 |
ndepth :: ('a, 'b) node => nat |
923 | 34 |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
35 |
Atom :: "('a + nat) => ('a, 'b) dtree" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
36 |
Leaf :: 'a => ('a, 'b) dtree |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
37 |
Numb :: nat => ('a, 'b) dtree |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
38 |
Scons :: [('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
39 |
In0,In1 :: ('a, 'b) dtree => ('a, 'b) dtree |
923 | 40 |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
41 |
Lim :: ('b => ('a, 'b) dtree) => ('a, 'b) dtree |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
42 |
Funs :: "'u set => ('t => 'u) set" |
923 | 43 |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
44 |
ntrunc :: [nat, ('a, 'b) dtree] => ('a, 'b) dtree |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
45 |
|
7255 | 46 |
uprod :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set |
47 |
usum :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set |
|
923 | 48 |
|
7255 | 49 |
Split :: [[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c |
50 |
Case :: [[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c |
|
923 | 51 |
|
7255 | 52 |
dprod :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] |
53 |
=> (('a, 'b) dtree * ('a, 'b) dtree)set" |
|
54 |
dsum :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] |
|
55 |
=> (('a, 'b) dtree * ('a, 'b) dtree)set" |
|
923 | 56 |
|
3947 | 57 |
|
58 |
local |
|
59 |
||
923 | 60 |
defs |
61 |
||
62 |
Push_Node_def "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))" |
|
63 |
||
64 |
(*crude "lists" of nats -- needed for the constructions*) |
|
1396 | 65 |
apfst_def "apfst == (%f (x,y). (f(x),y))" |
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
66 |
Push_def "Push == (%b h. nat_case b h)" |
923 | 67 |
|
68 |
(** operations on S-expressions -- sets of nodes **) |
|
69 |
||
70 |
(*S-expression constructors*) |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
71 |
Atom_def "Atom == (%x. {Abs_Node((%k. Inr 0, x))})" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
72 |
Scons_def "Scons M N == (Push_Node (Inr 1) `` M) Un (Push_Node (Inr 2) `` N)" |
923 | 73 |
|
74 |
(*Leaf nodes, with arbitrary or nat labels*) |
|
75 |
Leaf_def "Leaf == Atom o Inl" |
|
76 |
Numb_def "Numb == Atom o Inr" |
|
77 |
||
78 |
(*Injections of the "disjoint sum"*) |
|
5191
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents:
3947
diff
changeset
|
79 |
In0_def "In0(M) == Scons (Numb 0) M" |
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents:
3947
diff
changeset
|
80 |
In1_def "In1(M) == Scons (Numb 1) M" |
923 | 81 |
|
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
82 |
(*Function spaces*) |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
83 |
Lim_def "Lim f == Union {z. ? x. z = Push_Node (Inl x) `` (f x)}" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
84 |
Funs_def "Funs S == {f. range f <= S}" |
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
85 |
|
923 | 86 |
(*the set of nodes with depth less than k*) |
7014
11ee650edcd2
Added some definitions and theorems needed for the
berghofe
parents:
5978
diff
changeset
|
87 |
ndepth_def "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)" |
923 | 88 |
ntrunc_def "ntrunc k N == {n. n:N & ndepth(n)<k}" |
89 |
||
90 |
(*products and sums for the "universe"*) |
|
7255 | 91 |
uprod_def "uprod A B == UN x:A. UN y:B. { Scons x y }" |
92 |
usum_def "usum A B == In0``A Un In1``B" |
|
923 | 93 |
|
94 |
(*the corresponding eliminators*) |
|
5191
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents:
3947
diff
changeset
|
95 |
Split_def "Split c M == @u. ? x y. M = Scons x y & u = c x y" |
923 | 96 |
|
1151 | 97 |
Case_def "Case c d M == @u. (? x . M = In0(x) & u = c(x)) |
7255 | 98 |
| (? y . M = In1(y) & u = d(y))" |
923 | 99 |
|
100 |
||
5978
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
paulson
parents:
5191
diff
changeset
|
101 |
(** equality for the "universe" **) |
923 | 102 |
|
7255 | 103 |
dprod_def "dprod r s == UN (x,x'):r. UN (y,y'):s. {(Scons x y, Scons x' y')}" |
923 | 104 |
|
7255 | 105 |
dsum_def "dsum r s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un |
106 |
(UN (y,y'):s. {(In1(y),In1(y'))})" |
|
923 | 107 |
|
108 |
end |