Sum.thy
changeset 117 3716c99fb6a1
parent 107 960e332d2e70
child 122 6927e1cb2c07
--- a/Sum.thy	Fri Aug 19 11:10:56 1994 +0200
+++ b/Sum.thy	Fri Aug 19 11:15:01 1994 +0200
@@ -23,6 +23,9 @@
   Inl	   :: "'a => 'a+'b"
   Inr	   :: "'b => 'a+'b"
   sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c"
+
+  (*disjoint sum for sets; the operator + is overloaded with wrong type!*)
+  "plus"   :: "['a set, 'b set] => ('a+'b)set"      	(infixr 65)
   Part     :: "['a set, 'a=>'a] => 'a set"
 
 translations
@@ -32,18 +35,20 @@
   Inl_Rep_def	"Inl_Rep == (%a. %x y p. x=a & p)"
   Inr_Rep_def	"Inr_Rep == (%b. %x y p. y=b & ~p)"
 
-  Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
+  Sum_def 	"Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
     (*faking a type definition...*)
   Rep_Sum 		"Rep_Sum(s): Sum"
   Rep_Sum_inverse 	"Abs_Sum(Rep_Sum(s)) = s"
   Abs_Sum_inverse 	"f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
 
     (*defining the abstract constants*)
-  Inl_def  		"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
-  Inr_def 		"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
+  Inl_def  	"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
+  Inr_def 	"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
   sum_case_def	"sum_case(f,g,p) == @z.  (!x. p=Inl(x) --> z=f(x))	\
 \                                      & (!y. p=Inr(y) --> z=g(y))"
 
+  sum_def       "A plus B == (Inl``A) Un (Inr``B)"
+
   (*for selecting out the components of a mutually recursive definition*)
   Part_def	"Part(A,h) == A Int {x. ? z. x = h(z)}"