src/ZF/AC/HH.thy
author paulson
Mon, 21 May 2001 14:45:52 +0200
changeset 11317 7f9e4c389318
parent 1478 2b8c2a7547ab
child 12776 249600a63ba9
permissions -rw-r--r--
X-symbols for set theory

(*  Title:      ZF/AC/HH.thy
    ID:         $Id$
    Author:     Krzysztof Grabczewski

The theory file for the proofs of
  AC17 ==> AC1
  AC1 ==> WO2
  AC15 ==> WO6
*)

HH = AC_Equiv + Hartog + WO_AC + Let +

consts
 
  HH                      :: [i, i, i] => i

defs

  HH_def  "HH(f,x,a) == transrec(a, %b r. let z = x - (\\<Union>c \\<in> b. r`c)
                        in  if(f`z \\<in> Pow(z)-{0}, f`z, {x}))"

end