Sum.thy
author clasohm
Tue, 24 Oct 1995 14:59:17 +0100
changeset 251 f04b33ce250f
parent 249 492493334e0f
permissions -rw-r--r--
added calls of init_html and make_chart
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
     1
(*  Title:      HOL/Sum.thy
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
     6
The disjoint sum of two types.
0
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
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    11
(* type definition *)
51
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 38
diff changeset
    12
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    13
consts
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    14
  Inl_Rep       :: "['a, 'a, 'b, bool] => bool"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    15
  Inr_Rep       :: "['b, 'a, 'b, bool] => bool"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    16
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    17
defs
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    18
  Inl_Rep_def   "Inl_Rep == (%a. %x y p. x=a & p)"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    19
  Inr_Rep_def   "Inr_Rep == (%b. %x y p. y=b & ~p)"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    20
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    21
subtype (Sum)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    22
  ('a, 'b) "+"          (infixr 10)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    23
    = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    24
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    25
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    26
(* abstract constants and syntax *)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    27
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    28
consts
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    29
  Inl           :: "'a => 'a + 'b"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    30
  Inr           :: "'b => 'a + 'b"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    31
  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
    32
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    33
  (*disjoint sum for sets; the operator + is overloaded with wrong type!*)
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    34
  "plus"        :: "['a set, 'b set] => ('a + 'b) set"        (infixr 65)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    35
  Part          :: "['a set, 'b => 'a] => 'a set"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    36
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    37
translations
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    38
  "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
    39
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    40
defs
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    41
  Inl_def       "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    42
  Inr_def       "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
249
492493334e0f removed \...\ inside strings
clasohm
parents: 185
diff changeset
    43
  sum_case_def  "sum_case(f, g, p) == @z.  (!x. p=Inl(x) --> z=f(x))      
492493334e0f removed \...\ inside strings
clasohm
parents: 185
diff changeset
    44
                                        & (!y. p=Inr(y) --> z=g(y))"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    45
117
3716c99fb6a1 HOL/Sum: added disjoint sum of two sets as the plus operator, since + is
lcp
parents: 107
diff changeset
    46
  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
    47
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    48
  (*for selecting out the components of a mutually recursive definition*)
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 128
diff changeset
    49
  Part_def      "Part(A, h) == A Int {x. ? z. x = h(z)}"
107
960e332d2e70 HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents: 51
diff changeset
    50
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    51
end