:: Zermelo Theorem and Axiom of Choice. The correspondence of well ordering relations and ordinal numbers
:: by Grzegorz Bancerek
::
:: Received June 26, 1989
:: Copyright (c) 1990-2011 Association of Mizar Users


begin

definition
let X be set ;
func RelIncl X -> Relation means :Def1: :: WELLORD2:def 1
( field it = X & ( for Y, Z being set st Y in X & Z in X holds
( [Y,Z] in it iff Y c= Z ) ) );
existence
ex b1 being Relation st
( field b1 = X & ( for Y, Z being set st Y in X & Z in X holds
( [Y,Z] in b1 iff Y c= Z ) ) )
proof end;
uniqueness
for b1, b2 being Relation st field b1 = X & ( for Y, Z being set st Y in X & Z in X holds
( [Y,Z] in b1 iff Y c= Z ) ) & field b2 = X & ( for Y, Z being set st Y in X & Z in X holds
( [Y,Z] in b2 iff Y c= Z ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def1 defines RelIncl WELLORD2:def 1 :
for X being set
for b2 being Relation holds
( b2 = RelIncl X iff ( field b2 = X & ( for Y, Z being set st Y in X & Z in X holds
( [Y,Z] in b2 iff Y c= Z ) ) ) );

theorem :: WELLORD2:1
canceled;

theorem Th2: :: WELLORD2:2
for X being set holds RelIncl X is reflexive
proof end;

theorem Th3: :: WELLORD2:3
for X being set holds RelIncl X is transitive
proof end;

theorem Th4: :: WELLORD2:4
for A being Ordinal holds RelIncl A is connected
proof end;

theorem Th5: :: WELLORD2:5
for X being set holds RelIncl X is antisymmetric
proof end;

theorem Th6: :: WELLORD2:6
for A being Ordinal holds RelIncl A is well_founded
proof end;

theorem Th7: :: WELLORD2:7
for A being Ordinal holds RelIncl A is well-ordering
proof end;

theorem Th8: :: WELLORD2:8
for Y, X being set st Y c= X holds
(RelIncl X) |_2 Y = RelIncl Y
proof end;

theorem Th9: :: WELLORD2:9
for A being Ordinal
for X being set st X c= A holds
RelIncl X is well-ordering
proof end;

theorem Th10: :: WELLORD2:10
for A, B being Ordinal st A in B holds
A = (RelIncl B) -Seg A
proof end;

theorem Th11: :: WELLORD2:11
for A, B being Ordinal st RelIncl A, RelIncl B are_isomorphic holds
A = B
proof end;

theorem Th12: :: WELLORD2:12
for R being Relation
for A, B being Ordinal st R, RelIncl A are_isomorphic & R, RelIncl B are_isomorphic holds
A = B
proof end;

theorem Th13: :: WELLORD2:13
for R being Relation st R is well-ordering & ( for a being set st a in field R holds
ex A being Ordinal st R |_2 (R -Seg a), RelIncl A are_isomorphic ) holds
ex A being Ordinal st R, RelIncl A are_isomorphic
proof end;

theorem Th14: :: WELLORD2:14
for R being Relation st R is well-ordering holds
ex A being Ordinal st R, RelIncl A are_isomorphic
proof end;

definition
let R be Relation;
assume A1: R is well-ordering ;
func order_type_of R -> Ordinal means :Def2: :: WELLORD2:def 2
R, RelIncl it are_isomorphic ;
existence
ex b1 being Ordinal st R, RelIncl b1 are_isomorphic
by A1, Th14;
uniqueness
for b1, b2 being Ordinal st R, RelIncl b1 are_isomorphic & R, RelIncl b2 are_isomorphic holds
b1 = b2
by Th12;
end;

:: deftheorem Def2 defines order_type_of WELLORD2:def 2 :
for R being Relation st R is well-ordering holds
for b2 being Ordinal holds
( b2 = order_type_of R iff R, RelIncl b2 are_isomorphic );

definition
let A be Ordinal;
let R be Relation;
pred A is_order_type_of R means :: WELLORD2:def 3
A = order_type_of R;
end;

:: deftheorem defines is_order_type_of WELLORD2:def 3 :
for A being Ordinal
for R being Relation holds
( A is_order_type_of R iff A = order_type_of R );

theorem :: WELLORD2:15
canceled;

theorem :: WELLORD2:16
canceled;

theorem :: WELLORD2:17
for X being set
for A being Ordinal st X c= A holds
order_type_of (RelIncl X) c= A
proof end;

definition
let X, Y be set ;
:: original: are_equipotent
redefine pred X,Y are_equipotent means :: WELLORD2:def 4
ex f being Function st
( f is one-to-one & dom f = X & rng f = Y );
compatibility
( X,Y are_equipotent iff ex f being Function st
( f is one-to-one & dom f = X & rng f = Y ) )
proof end;
reflexivity
for X being set ex f being Function st
( f is one-to-one & dom f = X & rng f = X )
proof end;
symmetry
for X, Y being set st ex f being Function st
( f is one-to-one & dom f = X & rng f = Y ) holds
ex f being Function st
( f is one-to-one & dom f = Y & rng f = X )
proof end;
end;

:: deftheorem defines are_equipotent WELLORD2:def 4 :
for X, Y being set holds
( X,Y are_equipotent iff ex f being Function st
( f is one-to-one & dom f = X & rng f = Y ) );

theorem :: WELLORD2:18
canceled;

theorem :: WELLORD2:19
canceled;

theorem :: WELLORD2:20
canceled;

theorem :: WELLORD2:21
canceled;

theorem :: WELLORD2:22
for X, Y, Z being set st X,Y are_equipotent & Y,Z are_equipotent holds
X,Z are_equipotent
proof end;

theorem :: WELLORD2:23
canceled;

theorem :: WELLORD2:24
canceled;

theorem Th25: :: WELLORD2:25
for X being set
for R being Relation st R well_orders X holds
( field (R |_2 X) = X & R |_2 X is well-ordering )
proof end;

Lm1: for X being set
for R being Relation st R is well-ordering & X, field R are_equipotent holds
ex R being Relation st R well_orders X
proof end;

theorem Th26: :: WELLORD2:26
for X being set ex R being Relation st R well_orders X
proof end;

theorem :: WELLORD2:27
for M being non empty set st ( for X being set st X in M holds
X <> {} ) & ( for X, Y being set st X in M & Y in M & X <> Y holds
X misses Y ) holds
ex Choice being set st
for X being set st X in M holds
ex x being set st Choice /\ X = {x}
proof end;

theorem Th28: :: WELLORD2:28
for M being non empty set st ( for X being set st X in M holds
X <> {} ) holds
ex Choice being Function st
( dom Choice = M & ( for X being set st X in M holds
Choice . X in X ) )
proof end;

begin

scheme :: WELLORD2:sch 1
NonUniqBoundFuncEx{ F1() -> set , F2() -> set , P1[ set , set ] } :
ex f being Function st
( dom f = F1() & rng f c= F2() & ( for x being set st x in F1() holds
P1[x,f . x] ) )
provided
A1: for x being set st x in F1() holds
ex y being set st
( y in F2() & P1[x,y] )
proof end;

scheme :: WELLORD2:sch 2
BiFuncEx{ F1() -> set , F2() -> set , F3() -> set , P1[ set , set , set ] } :
ex f, g being Function st
( dom f = F1() & dom g = F1() & ( for x being set st x in F1() holds
P1[x,f . x,g . x] ) )
provided
A1: for x being set st x in F1() holds
ex y, z being set st
( y in F2() & z in F3() & P1[x,y,z] )
proof end;

theorem :: WELLORD2:29
for X being set holds RelIncl X is_reflexive_in X
proof end;

theorem :: WELLORD2:30
for X being set holds RelIncl X is_transitive_in X
proof end;

theorem :: WELLORD2:31
for X being set holds RelIncl X is_antisymmetric_in X
proof end;

registration
cluster RelIncl {} -> empty ;
coherence
RelIncl {} is empty
proof end;
end;

registration
let X be non empty set ;
cluster RelIncl X -> non empty ;
coherence
not RelIncl X is empty
proof end;
end;

theorem :: WELLORD2:32
for x being set holds RelIncl {x} = {[x,x]}
proof end;

theorem :: WELLORD2:33
for X being set holds RelIncl X c= [:X,X:]
proof end;