(* Title: HOL/sum
ID: $Id$
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1992 University of Cambridge
The disjoint sum of two types
*)
Sum = Prod +
types
('a,'b) "+" (infixr 10)
arities
"+" :: (term,term)term
consts
Inl_Rep :: "['a,'a,'b,bool] => bool"
Inr_Rep :: "['b,'a,'b,bool] => bool"
Sum :: "(['a,'b,bool] => bool)set"
Rep_Sum :: "'a + 'b => (['a,'b,bool] => bool)"
Abs_Sum :: "(['a,'b,bool] => bool) => 'a+'b"
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, 'b=>'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(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)}"
end