| author | nipkow | 
| Thu, 06 May 1999 11:13:01 +0200 | |
| changeset 6605 | c2754409919b | 
| parent 6053 | 8a1059aa01f0 | 
| child 13134 | bf37a3049251 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/Zorn.thy | 
| 516 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 516 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 6 | Based upon the article | |
| 7 | Abrial & Laffitte, | |
| 8 | Towards the Mechanization of the Proofs of Some | |
| 9 | Classical Theorems of Set Theory. | |
| 10 | ||
| 11 | Union_in_Pow is proved in ZF.ML | |
| 12 | *) | |
| 13 | ||
| 806 | 14 | Zorn = OrderArith + AC + Inductive + | 
| 516 | 15 | |
| 16 | consts | |
| 1401 | 17 | Subset_rel :: i=>i | 
| 18 | chain, maxchain :: i=>i | |
| 19 | super :: [i,i]=>i | |
| 516 | 20 | |
| 753 | 21 | defs | 
| 516 | 22 |   Subset_rel_def "Subset_rel(A) == {z: A*A . EX x y. z=<x,y> & x<=y & x~=y}"
 | 
| 485 | 23 | |
| 516 | 24 |   chain_def      "chain(A)      == {F: Pow(A). ALL X:F. ALL Y:F. X<=Y | Y<=X}"
 | 
| 25 |   super_def      "super(A,c)    == {d: chain(A). c<=d & c~=d}"
 | |
| 26 |   maxchain_def   "maxchain(A)   == {c: chain(A). super(A,c)=0}"
 | |
| 27 | ||
| 28 | ||
| 29 | (** We could make the inductive definition conditional on next: increasing(S) | |
| 30 | but instead we make this a side-condition of an introduction rule. Thus | |
| 31 | the induction rule lets us assume that condition! Many inductive proofs | |
| 32 | are therefore unconditional. | |
| 33 | **) | |
| 34 | consts | |
| 1401 | 35 | "TFin" :: [i,i]=>i | 
| 516 | 36 | |
| 37 | inductive | |
| 38 | domains "TFin(S,next)" <= "Pow(S)" | |
| 39 | intrs | |
| 1478 | 40 | nextI "[| x : TFin(S,next); next: increasing(S) | 
| 1155 | 41 | |] ==> next`x : TFin(S,next)" | 
| 516 | 42 | |
| 1478 | 43 | Pow_UnionI "Y : Pow(TFin(S,next)) ==> Union(Y) : TFin(S,next)" | 
| 516 | 44 | |
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1478diff
changeset | 45 | monos Pow_mono | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1478diff
changeset | 46 | con_defs increasing_def | 
| 516 | 47 | type_intrs "[CollectD1 RS apply_funtype, Union_in_Pow]" | 
| 48 | ||
| 49 | end |