src/ZF/Sum.thy
author paulson
Mon, 17 Aug 1998 13:09:08 +0200
changeset 5325 f7a5e06adea1
parent 3940 1d5bee4d047f
child 13220 62c899c77151
permissions -rw-r--r--
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy contains fewer theorems than before
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/sum.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Disjoint sums in Zermelo-Fraenkel Set Theory 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
"Part" primitive for simultaneous recursive type definitions
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 3940
diff changeset
    10
Sum = Bool + equalities + 
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    11
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    12
global
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    13
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
consts
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    15
    "+"         :: [i,i]=>i                     (infixr 65)
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1108
diff changeset
    16
    Inl,Inr     :: i=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1108
diff changeset
    17
    case        :: [i=>i, i=>i, i]=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1108
diff changeset
    18
    Part        :: [i,i=>i] => i
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
3940
wenzelm
parents: 3923
diff changeset
    20
local
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    21
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 124
diff changeset
    22
defs
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
    sum_def     "A+B == {0}*A Un {1}*B"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
    Inl_def     "Inl(a) == <0,a>"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
    Inr_def     "Inr(b) == <1,b>"
1108
22b256c8c9fb case is defined using pattern-matching
lcp
parents: 753
diff changeset
    26
    case_def    "case(c,d) == (%<y,z>. cond(y, d(z), c(z)))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
  (*operator for selecting out the various summands*)
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    29
    Part_def    "Part(A,h) == {x: A. EX z. x = h(z)}"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
end