author | nipkow |
Mon, 17 Aug 1998 11:00:27 +0200 | |
changeset 5322 | 504b129e0502 |
parent 3947 | eb707467f8c5 |
child 7254 | fc7f95f293da |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/Sum.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
The disjoint sum of two types. |
|
7 |
*) |
|
8 |
||
1515 | 9 |
Sum = mono + Prod + |
923 | 10 |
|
11 |
(* type definition *) |
|
12 |
||
1558 | 13 |
constdefs |
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
14 |
Inl_Rep :: ['a, 'a, 'b, bool] => bool |
1558 | 15 |
"Inl_Rep == (%a. %x y p. x=a & p)" |
923 | 16 |
|
1558 | 17 |
Inr_Rep :: ['b, 'a, 'b, bool] => bool |
18 |
"Inr_Rep == (%b. %x y p. y=b & ~p)" |
|
923 | 19 |
|
3947 | 20 |
global |
21 |
||
1475 | 22 |
typedef (Sum) |
923 | 23 |
('a, 'b) "+" (infixr 10) |
24 |
= "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}" |
|
25 |
||
26 |
||
27 |
(* abstract constants and syntax *) |
|
28 |
||
29 |
consts |
|
30 |
Inl :: "'a => 'a + 'b" |
|
31 |
Inr :: "'b => 'a + 'b" |
|
32 |
sum_case :: "['a => 'c, 'b => 'c, 'a + 'b] => 'c" |
|
33 |
||
34 |
(*disjoint sum for sets; the operator + is overloaded with wrong type!*) |
|
2212 | 35 |
Plus :: "['a set, 'b set] => ('a + 'b) set" (infixr 65) |
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
36 |
Part :: ['a set, 'b => 'a] => 'a set |
923 | 37 |
|
38 |
translations |
|
3842 | 39 |
"case p of Inl x => a | Inr y => b" == "sum_case (%x. a) (%y. b) p" |
923 | 40 |
|
3947 | 41 |
local |
42 |
||
923 | 43 |
defs |
44 |
Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))" |
|
45 |
Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))" |
|
1151 | 46 |
sum_case_def "sum_case f g p == @z. (!x. p=Inl(x) --> z=f(x)) |
1423 | 47 |
& (!y. p=Inr(y) --> z=g(y))" |
923 | 48 |
|
2212 | 49 |
sum_def "A Plus B == (Inl``A) Un (Inr``B)" |
923 | 50 |
|
51 |
(*for selecting out the components of a mutually recursive definition*) |
|
52 |
Part_def "Part A h == A Int {x. ? z. x = h(z)}" |
|
53 |
||
54 |
end |