| author | paulson | 
| Tue, 22 Sep 1998 15:24:39 +0200 | |
| changeset 5533 | bce36a019b03 | 
| parent 2469 | b50b8c0eec01 | 
| child 6053 | 8a1059aa01f0 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/AC.thy | 
| 484 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 484 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 6 | The Axiom of Choice | |
| 7 | ||
| 8 | This definition comes from Halmos (1960), page 59. | |
| 9 | *) | |
| 10 | ||
| 2469 | 11 | AC = func + | 
| 484 | 12 | rules | 
| 1478 | 13 | AC "[| a: A; !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)" | 
| 484 | 14 | end |