Sum.thy
author lcp
Fri, 19 Aug 1994 11:15:01 +0200
changeset 117 3716c99fb6a1
parent 107 960e332d2e70
child 122 6927e1cb2c07
permissions -rw-r--r--
HOL/Sum: added disjoint sum of two sets as the plus operator, since + is overloaded with an incompatible type
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/sum
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     6
The disjoint sum of two types
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     7
*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     8
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     9
Sum = Prod +
51
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    10
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    11
types
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    12
  ('a,'b) "+"		      (infixl 10)
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    13
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    14
arities
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    15
  "+"      :: (term,term)term
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    16
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    17
consts
51
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    18
  Inl_Rep  :: "['a,'a,'b,bool] => bool"
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    19
  Inr_Rep  :: "['b,'a,'b,bool] => bool"
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    20
  Sum      :: "(['a,'b,bool] => bool)set"
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    21
  Rep_Sum  :: "'a + 'b => (['a,'b,bool] => bool)"
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    22
  Abs_Sum  :: "(['a,'b,bool] => bool) => 'a+'b"
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    23
  Inl	   :: "'a => 'a+'b"
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    24
  Inr	   :: "'b => 'a+'b"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    25
  sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c"
117
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    26
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    27
  (*disjoint sum for sets; the operator + is overloaded with wrong type!*)
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    28
  "plus"   :: "['a set, 'b set] => ('a+'b)set"      	(infixr 65)
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    29
  Part     :: "['a set, 'a=>'a] => 'a set"
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    30
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    31
translations
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    32
  "case p of Inl(x) => a | Inr(y) => b" == "sum_case(%x.a, %y.b, p)"
51
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    33
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    34
rules
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    35
  Inl_Rep_def	"Inl_Rep == (%a. %x y p. x=a & p)"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    36
  Inr_Rep_def	"Inr_Rep == (%b. %x y p. y=b & ~p)"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    37
117
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    38
  Sum_def 	"Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    39
    (*faking a type definition...*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    40
  Rep_Sum 		"Rep_Sum(s): Sum"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    41
  Rep_Sum_inverse 	"Abs_Sum(Rep_Sum(s)) = s"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    42
  Abs_Sum_inverse 	"f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    43
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    44
    (*defining the abstract constants*)
117
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    45
  Inl_def  	"Inl == (%a. Abs_Sum(Inl_Rep(a)))"
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    46
  Inr_def 	"Inr == (%b. Abs_Sum(Inr_Rep(b)))"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    47
  sum_case_def	"sum_case(f,g,p) == @z.  (!x. p=Inl(x) --> z=f(x))	\
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    48
\                                      & (!y. p=Inr(y) --> z=g(y))"
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    49
117
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    50
  sum_def       "A plus B == (Inl``A) Un (Inr``B)"
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    51
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    52
  (*for selecting out the components of a mutually recursive definition*)
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    53
  Part_def	"Part(A,h) == A Int {x. ? z. x = h(z)}"
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    54
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    55
end