diff -r 000000000000 -r 7949f97df77a Sum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Sum.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,34 @@ +(* 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 "+" 2 (infixl 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" + case :: "['a+'b, 'a=>'c,'b=>'c] =>'c" +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)))" + case_def "case == (%p f g. @z. (!x. p=Inl(x) --> z=f(x))\ +\ & (!y. p=Inr(y) --> z=g(y)))" +end