author | clasohm |
Fri, 19 Apr 1996 11:33:24 +0200 | |
changeset 1668 | 8ead1fe65aad |
parent 1478 | 2b8c2a7547ab |
child 2469 | b50b8c0eec01 |
permissions | -rw-r--r-- |
(* Title: ZF/AC.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge The Axiom of Choice This definition comes from Halmos (1960), page 59. *) AC = ZF + "func" + rules AC "[| a: A; !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)" end