Sum.thy
changeset 107 960e332d2e70
parent 51 934a58983311
child 117 3716c99fb6a1
equal deleted inserted replaced
106:d27056ec0a5a 107:960e332d2e70
    20   Sum      :: "(['a,'b,bool] => bool)set"
    20   Sum      :: "(['a,'b,bool] => bool)set"
    21   Rep_Sum  :: "'a + 'b => (['a,'b,bool] => bool)"
    21   Rep_Sum  :: "'a + 'b => (['a,'b,bool] => bool)"
    22   Abs_Sum  :: "(['a,'b,bool] => bool) => 'a+'b"
    22   Abs_Sum  :: "(['a,'b,bool] => bool) => 'a+'b"
    23   Inl	   :: "'a => 'a+'b"
    23   Inl	   :: "'a => 'a+'b"
    24   Inr	   :: "'b => 'a+'b"
    24   Inr	   :: "'b => 'a+'b"
    25   sum_case :: "['a+'b, 'a=>'c,'b=>'c] =>'c"
    25   sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c"
       
    26   Part     :: "['a set, 'a=>'a] => 'a set"
       
    27 
       
    28 translations
       
    29   "case p of Inl(x) => a | Inr(y) => b" == "sum_case(%x.a, %y.b, p)"
    26 
    30 
    27 rules
    31 rules
    28   Inl_Rep_def	"Inl_Rep == (%a. %x y p. x=a & p)"
    32   Inl_Rep_def	"Inl_Rep == (%a. %x y p. x=a & p)"
    29   Inr_Rep_def	"Inr_Rep == (%b. %x y p. y=b & ~p)"
    33   Inr_Rep_def	"Inr_Rep == (%b. %x y p. y=b & ~p)"
       
    34 
    30   Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
    35   Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
    31     (*faking a type definition...*)
    36     (*faking a type definition...*)
    32   Rep_Sum 		"Rep_Sum(s): Sum"
    37   Rep_Sum 		"Rep_Sum(s): Sum"
    33   Rep_Sum_inverse 	"Abs_Sum(Rep_Sum(s)) = s"
    38   Rep_Sum_inverse 	"Abs_Sum(Rep_Sum(s)) = s"
    34   Abs_Sum_inverse 	"f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
    39   Abs_Sum_inverse 	"f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
       
    40 
    35     (*defining the abstract constants*)
    41     (*defining the abstract constants*)
    36   Inl_def  		"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
    42   Inl_def  		"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
    37   Inr_def 		"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
    43   Inr_def 		"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
    38   sum_case_def	"sum_case == (%p f g. @z.  (!x. p=Inl(x) --> z=f(x))\
    44   sum_case_def	"sum_case(f,g,p) == @z.  (!x. p=Inl(x) --> z=f(x))	\
    39 \                                        & (!y. p=Inr(y) --> z=g(y)))"
    45 \                                      & (!y. p=Inr(y) --> z=g(y))"
       
    46 
       
    47   (*for selecting out the components of a mutually recursive definition*)
       
    48   Part_def	"Part(A,h) == A Int {x. ? z. x = h(z)}"
       
    49 
    40 end
    50 end