diff -r d27056ec0a5a -r 960e332d2e70 Sum.thy --- a/Sum.thy Thu Aug 18 11:40:54 1994 +0200 +++ b/Sum.thy Thu Aug 18 11:43:40 1994 +0200 @@ -22,19 +22,29 @@ Abs_Sum :: "(['a,'b,bool] => bool) => 'a+'b" Inl :: "'a => 'a+'b" Inr :: "'b => 'a+'b" - sum_case :: "['a+'b, 'a=>'c,'b=>'c] =>'c" + sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c" + Part :: "['a set, 'a=>'a] => 'a set" + +translations + "case p of Inl(x) => a | Inr(y) => b" == "sum_case(%x.a, %y.b, p)" rules 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))}" (*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)))" - sum_case_def "sum_case == (%p f g. @z. (!x. p=Inl(x) --> z=f(x))\ -\ & (!y. p=Inr(y) --> z=g(y)))" + sum_case_def "sum_case(f,g,p) == @z. (!x. p=Inl(x) --> z=f(x)) \ +\ & (!y. p=Inr(y) --> z=g(y))" + + (*for selecting out the components of a mutually recursive definition*) + Part_def "Part(A,h) == A Int {x. ? z. x = h(z)}" + end