src/ZF/Zorn0.thy
changeset 485 5e00a676a211
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/Zorn0.thy	Tue Jul 26 13:44:42 1994 +0200
@@ -0,0 +1,28 @@
+(*  Title: 	ZF/Zorn0.thy
+    ID:         $Id$
+    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1994  University of Cambridge
+
+Based upon the article
+    Abrial & Laffitte, 
+    Towards the Mechanization of the Proofs of Some 
+    Classical Theorems of Set Theory. 
+*)
+
+Zorn0 = OrderArith + AC + "inductive" +
+
+consts
+  Subset_rel      :: "i=>i"
+  increasing      :: "i=>i"
+  chain, maxchain :: "i=>i"
+  super           :: "[i,i]=>i"
+
+rules
+  Subset_rel_def "Subset_rel(A) == {z: A*A . EX x y. z=<x,y> & x<=y & x~=y}"
+  increasing_def "increasing(A) == {f: Pow(A)->Pow(A). ALL x. x<=A --> x<=f`x}"
+
+  chain_def      "chain(A)      == {F: Pow(A). ALL X:F. ALL Y:F. X<=Y | Y<=X}"
+  super_def      "super(A,c)    == {d: chain(A). c<=d & c~=d}"
+  maxchain_def   "maxchain(A)   == {c: chain(A). super(A,c)=0}"
+
+end