src/ZF/Sum.thy
author wenzelm
Thu, 08 Nov 2001 23:52:56 +0100
changeset 12106 4a8558dbb6a0
parent 5325 f7a5e06adea1
child 13220 62c899c77151
permissions -rw-r--r--
* Isar/Pure: emulation of instantiation tactics (rule_tac, cut_tac, etc.) now consider the syntactic context of assumptions, giving a better chance to get type-inference of the arguments right (this is especially important for locales); * system: refrain from any attempt at filtering input streams; no longer support ``8bit'' encoding of old isabelle font, instead proper iso-latin characters may now be used;

(*  Title:      ZF/sum.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge

Disjoint sums in Zermelo-Fraenkel Set Theory 
"Part" primitive for simultaneous recursive type definitions
*)

Sum = Bool + equalities + 

global

consts
    "+"         :: [i,i]=>i                     (infixr 65)
    Inl,Inr     :: i=>i
    case        :: [i=>i, i=>i, i]=>i
    Part        :: [i,i=>i] => i

local

defs
    sum_def     "A+B == {0}*A Un {1}*B"
    Inl_def     "Inl(a) == <0,a>"
    Inr_def     "Inr(b) == <1,b>"
    case_def    "case(c,d) == (%<y,z>. cond(y, d(z), c(z)))"

  (*operator for selecting out the various summands*)
    Part_def    "Part(A,h) == {x: A. EX z. x = h(z)}"
end