Sum.thy
changeset 38 7ef6ba42914b
parent 0 7949f97df77a
child 51 934a58983311
equal deleted inserted replaced
37:65a546c2b8ed 38:7ef6ba42914b
    15    Sum		:: "(['a,'b,bool] => bool)set"
    15    Sum		:: "(['a,'b,bool] => bool)set"
    16    Rep_Sum	:: "'a + 'b => (['a,'b,bool] => bool)"
    16    Rep_Sum	:: "'a + 'b => (['a,'b,bool] => bool)"
    17    Abs_Sum	:: "(['a,'b,bool] => bool) => 'a+'b"
    17    Abs_Sum	:: "(['a,'b,bool] => bool) => 'a+'b"
    18    Inl		:: "'a => 'a+'b"
    18    Inl		:: "'a => 'a+'b"
    19    Inr		:: "'b => 'a+'b"
    19    Inr		:: "'b => 'a+'b"
    20    case		:: "['a+'b, 'a=>'c,'b=>'c] =>'c"
    20    sum_case	:: "['a+'b, 'a=>'c,'b=>'c] =>'c"
    21 rules
    21 rules
    22   Inl_Rep_def	"Inl_Rep == (%a. %x y p. x=a & p)"
    22   Inl_Rep_def	"Inl_Rep == (%a. %x y p. x=a & p)"
    23   Inr_Rep_def	"Inr_Rep == (%b. %x y p. y=b & ~p)"
    23   Inr_Rep_def	"Inr_Rep == (%b. %x y p. y=b & ~p)"
    24   Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
    24   Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
    25     (*faking a type definition...*)
    25     (*faking a type definition...*)
    27   Rep_Sum_inverse 	"Abs_Sum(Rep_Sum(s)) = s"
    27   Rep_Sum_inverse 	"Abs_Sum(Rep_Sum(s)) = s"
    28   Abs_Sum_inverse 	"f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
    28   Abs_Sum_inverse 	"f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
    29     (*defining the abstract constants*)
    29     (*defining the abstract constants*)
    30   Inl_def  		"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
    30   Inl_def  		"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
    31   Inr_def 		"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
    31   Inr_def 		"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
    32   case_def	"case == (%p f g. @z.  (!x. p=Inl(x) --> z=f(x))\
    32   sum_case_def	"sum_case == (%p f g. @z.  (!x. p=Inl(x) --> z=f(x))\
    33 \                                    & (!y. p=Inr(y) --> z=g(y)))"
    33 \                                        & (!y. p=Inr(y) --> z=g(y)))"
    34 end
    34 end