| author | paulson | 
| Sun, 23 Jun 2002 10:14:13 +0200 | |
| changeset 13240 | bb5f4faea1f3 | 
| parent 11609 | 3f3d1add4d94 | 
| child 15391 | 797ed46d724b | 
| permissions | -rw-r--r-- | 
| 10213 | 1  | 
(* Title: HOL/Sum_Type.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  | 
||
| 
11609
 
3f3d1add4d94
eliminated theories "equalities" and "mono" (made part of "Typedef",
 
wenzelm 
parents: 
10832 
diff
changeset
 | 
9  | 
Sum_Type = Product_Type +  | 
| 10213 | 10  | 
|
11  | 
(* type definition *)  | 
|
12  | 
||
13  | 
constdefs  | 
|
14  | 
Inl_Rep :: ['a, 'a, 'b, bool] => bool  | 
|
15  | 
"Inl_Rep == (%a. %x y p. x=a & p)"  | 
|
16  | 
||
17  | 
Inr_Rep :: ['b, 'a, 'b, bool] => bool  | 
|
18  | 
"Inr_Rep == (%b. %x y p. y=b & ~p)"  | 
|
19  | 
||
20  | 
global  | 
|
21  | 
||
22  | 
typedef (Sum)  | 
|
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  | 
||
33  | 
(*disjoint sum for sets; the operator + is overloaded with wrong type!*)  | 
|
34  | 
  Plus          :: "['a set, 'b set] => ('a + 'b) set"        (infixr "<+>" 65)
 | 
|
35  | 
Part :: ['a set, 'b => 'a] => 'a set  | 
|
36  | 
||
37  | 
local  | 
|
38  | 
||
39  | 
defs  | 
|
40  | 
Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))"  | 
|
41  | 
Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))"  | 
|
42  | 
||
| 10832 | 43  | 
sum_def "A <+> B == (Inl`A) Un (Inr`B)"  | 
| 10213 | 44  | 
|
45  | 
(*for selecting out the components of a mutually recursive definition*)  | 
|
46  | 
  Part_def      "Part A h == A Int {x. ? z. x = h(z)}"
 | 
|
47  | 
||
48  | 
end  |