:: ORDINAL1 semantic presentation

theorem Th1: :: ORDINAL1:1
canceled;

theorem Th2: :: ORDINAL1:2
canceled;

theorem Th3: :: ORDINAL1:3
for b1, b2, b3 being set holds
not ( b1 in b2 & b2 in b3 & b3 in b1 )
proof end;

theorem Th4: :: ORDINAL1:4
for b1, b2, b3, b4 being set holds
not ( b1 in b2 & b2 in b3 & b3 in b4 & b4 in b1 )
proof end;

theorem Th5: :: ORDINAL1:5
for b1, b2, b3, b4, b5 being set holds
not ( b1 in b2 & b2 in b3 & b3 in b4 & b4 in b5 & b5 in b1 )
proof end;

theorem Th6: :: ORDINAL1:6
for b1, b2, b3, b4, b5, b6 being set holds
not ( b1 in b2 & b2 in b3 & b3 in b4 & b4 in b5 & b5 in b6 & b6 in b1 )
proof end;

theorem Th7: :: ORDINAL1:7
for b1, b2 being set holds
not ( b1 in b2 & b2 c= b1 )
proof end;

definition
let c1 be set ;
func succ c1 -> set equals :: ORDINAL1:def 1
a1 \/ {a1};
coherence
c1 \/ {c1} is set
;
end;

:: deftheorem Def1 defines succ ORDINAL1:def 1 :
for b1 being set holds succ b1 = b1 \/ {b1};

registration
let c1 be set ;
cluster succ a1 -> non empty ;
coherence
not succ c1 is empty
;
end;

theorem Th8: :: ORDINAL1:8
canceled;

theorem Th9: :: ORDINAL1:9
canceled;

theorem Th10: :: ORDINAL1:10
for b1 being set holds b1 in succ b1
proof end;

theorem Th11: :: ORDINAL1:11
canceled;

theorem Th12: :: ORDINAL1:12
for b1, b2 being set holds
( succ b1 = succ b2 implies b1 = b2 )
proof end;

theorem Th13: :: ORDINAL1:13
for b1, b2 being set holds
( b1 in succ b2 iff ( b1 in b2 or b1 = b2 ) )
proof end;

theorem Th14: :: ORDINAL1:14
for b1 being set holds
b1 <> succ b1
proof end;

definition
let c1 be set ;
attr a1 is epsilon-transitive means :Def2: :: ORDINAL1:def 2
for b1 being set holds
( b1 in a1 implies b1 c= a1 );
attr a1 is epsilon-connected means :Def3: :: ORDINAL1:def 3
for b1, b2 being set holds
not ( b1 in a1 & b2 in a1 & not b1 in b2 & not b1 = b2 & not b2 in b1 );
end;

:: deftheorem Def2 defines epsilon-transitive ORDINAL1:def 2 :
for b1 being set holds
( b1 is epsilon-transitive iff for b2 being set holds
( b2 in b1 implies b2 c= b1 ) );

:: deftheorem Def3 defines epsilon-connected ORDINAL1:def 3 :
for b1 being set holds
( b1 is epsilon-connected iff for b2, b3 being set holds
not ( b2 in b1 & b3 in b1 & not b2 in b3 & not b2 = b3 & not b3 in b2 ) );

Lemma8: ( {} is epsilon-transitive & {} is epsilon-connected )
proof end;

definition
let c1 be set ;
attr a1 is ordinal means :Def4: :: ORDINAL1:def 4
( a1 is epsilon-transitive & a1 is epsilon-connected );
end;

:: deftheorem Def4 defines ordinal ORDINAL1:def 4 :
for b1 being set holds
( b1 is ordinal iff ( b1 is epsilon-transitive & b1 is epsilon-connected ) );

registration
cluster ordinal -> epsilon-transitive epsilon-connected number ;
coherence
for b1 being set holds
( b1 is ordinal implies ( b1 is epsilon-transitive & b1 is epsilon-connected ) )
by Def4;
cluster epsilon-transitive epsilon-connected -> ordinal number ;
coherence
for b1 being set holds
( b1 is epsilon-transitive & b1 is epsilon-connected implies b1 is ordinal )
by Def4;
end;

notation
synonym number for set ;
end;

registration
cluster epsilon-transitive epsilon-connected ordinal number ;
existence
ex b1 being number st b1 is ordinal
proof end;
end;

definition
mode Ordinal is ordinal number ;
end;

theorem Th15: :: ORDINAL1:15
canceled;

theorem Th16: :: ORDINAL1:16
canceled;

theorem Th17: :: ORDINAL1:17
canceled;

theorem Th18: :: ORDINAL1:18
canceled;

theorem Th19: :: ORDINAL1:19
for b1, b2 being Ordinal
for b3 being epsilon-transitive set holds
( b3 in b1 & b1 in b2 implies b3 in b2 )
proof end;

theorem Th20: :: ORDINAL1:20
canceled;

theorem Th21: :: ORDINAL1:21
for b1 being epsilon-transitive set
for b2 being Ordinal holds
( b1 c< b2 implies b1 in b2 )
proof end;

theorem Th22: :: ORDINAL1:22
for b1 being epsilon-transitive set
for b2, b3 being Ordinal holds
( b1 c= b2 & b2 in b3 implies b1 in b3 )
proof end;

theorem Th23: :: ORDINAL1:23
for b1 being set
for b2 being Ordinal holds
( b1 in b2 implies b1 is Ordinal )
proof end;

theorem Th24: :: ORDINAL1:24
for b1, b2 being Ordinal holds
not ( not b1 in b2 & not b1 = b2 & not b2 in b1 )
proof end;

definition
let c1, c2 be Ordinal;
redefine pred c= as c1 c= c2;
connectedness
for b1, b2 being Ordinal holds
( not b1 c= b2 implies b2 c= b1 )
proof end;
end;

theorem Th25: :: ORDINAL1:25
for b1, b2 being Ordinal holds b1,b2 are_c=-comparable
proof end;

theorem Th26: :: ORDINAL1:26
for b1, b2 being Ordinal holds
( b1 c= b2 or b2 in b1 )
proof end;

theorem Th27: :: ORDINAL1:27
{} is Ordinal by Def4, Lemma8;

registration
cluster empty number ;
existence
ex b1 being Ordinal st b1 is empty
by Th27;
end;

registration
cluster empty -> epsilon-transitive epsilon-connected ordinal number ;
coherence
for b1 being number holds
( b1 is empty implies b1 is ordinal )
by Def4, Lemma8;
end;

registration
cluster {} -> epsilon-transitive epsilon-connected ordinal ;
coherence
{} is ordinal
;
end;

theorem Th28: :: ORDINAL1:28
canceled;

theorem Th29: :: ORDINAL1:29
for b1 being set holds
( b1 is Ordinal implies succ b1 is Ordinal )
proof end;

theorem Th30: :: ORDINAL1:30
for b1 being set holds
( b1 is ordinal implies union b1 is ordinal )
proof end;

registration
cluster non empty number ;
existence
not for b1 being Ordinal holds b1 is empty
proof end;
end;

registration
let c1 be Ordinal;
cluster succ a1 -> non empty epsilon-transitive epsilon-connected ordinal ;
coherence
( not succ c1 is empty & succ c1 is ordinal )
by Th29;
cluster union a1 -> epsilon-transitive epsilon-connected ordinal ;
coherence
union c1 is ordinal
by Th30;
end;

theorem Th31: :: ORDINAL1:31
for b1 being set holds
( ( for b2 being set holds
( b2 in b1 implies ( b2 is Ordinal & b2 c= b1 ) ) ) implies b1 is ordinal )
proof end;

theorem Th32: :: ORDINAL1:32
for b1 being set
for b2 being Ordinal holds
not ( b1 c= b2 & b1 <> {} & ( for b3 being Ordinal holds
not ( b3 in b1 & ( for b4 being Ordinal holds
( b4 in b1 implies b3 c= b4 ) ) ) ) )
proof end;

theorem Th33: :: ORDINAL1:33
for b1, b2 being Ordinal holds
( b1 in b2 iff succ b1 c= b2 )
proof end;

theorem Th34: :: ORDINAL1:34
for b1, b2 being Ordinal holds
( b1 in succ b2 iff b1 c= b2 )
proof end;

scheme :: ORDINAL1:sch 1
s1{ P1[ Ordinal] } :
ex b1 being Ordinal st
( P1[b1] & ( for b2 being Ordinal holds
( P1[b2] implies b1 c= b2 ) ) )
provided
E22: ex b1 being Ordinal st P1[b1]
proof end;

scheme :: ORDINAL1:sch 2
s2{ P1[ Ordinal] } :
for b1 being Ordinal holds P1[b1]
provided
E22: for b1 being Ordinal holds
( ( for b2 being Ordinal holds
( b2 in b1 implies P1[b2] ) ) implies P1[b1] )
proof end;

theorem Th35: :: ORDINAL1:35
for b1 being set holds
( ( for b2 being set holds
( b2 in b1 implies b2 is Ordinal ) ) implies union b1 is ordinal )
proof end;

theorem Th36: :: ORDINAL1:36
for b1 being set holds
not ( ( for b2 being set holds
( b2 in b1 implies b2 is Ordinal ) ) & ( for b2 being Ordinal holds
not b1 c= b2 ) )
proof end;

theorem Th37: :: ORDINAL1:37
for b1 being set holds
not for b2 being set holds
( b2 in b1 iff b2 is Ordinal )
proof end;

theorem Th38: :: ORDINAL1:38
for b1 being set holds
not for b2 being Ordinal holds b2 in b1
proof end;

theorem Th39: :: ORDINAL1:39
for b1 being set holds
ex b2 being Ordinal st
( not b2 in b1 & ( for b3 being Ordinal holds
( not b3 in b1 implies b2 c= b3 ) ) )
proof end;

definition
let c1 be set ;
canceled;
attr a1 is being_limit_ordinal means :Def6: :: ORDINAL1:def 6
a1 = union a1;
end;

:: deftheorem Def5 ORDINAL1:def 5 :
canceled;

:: deftheorem Def6 defines being_limit_ordinal ORDINAL1:def 6 :
for b1 being set holds
( b1 is being_limit_ordinal iff b1 = union b1 );

notation
let c1 be set ;
synonym c1 is_limit_ordinal for being_limit_ordinal c1;
end;

theorem Th40: :: ORDINAL1:40
canceled;

theorem Th41: :: ORDINAL1:41
for b1 being Ordinal holds
( b1 is_limit_ordinal iff for b2 being Ordinal holds
( b2 in b1 implies succ b2 in b1 ) )
proof end;

theorem Th42: :: ORDINAL1:42
for b1 being Ordinal holds
( not ( not b1 is_limit_ordinal & ( for b2 being Ordinal holds
not b1 = succ b2 ) ) & not ( ex b2 being Ordinal st b1 = succ b2 & b1 is_limit_ordinal ) )
proof end;

definition
let c1 be Function;
attr a1 is T-Sequence-like means :Def7: :: ORDINAL1:def 7
dom a1 is ordinal;
end;

:: deftheorem Def7 defines T-Sequence-like ORDINAL1:def 7 :
for b1 being Function holds
( b1 is T-Sequence-like iff dom b1 is ordinal );

registration
cluster T-Sequence-like number ;
existence
ex b1 being Function st b1 is T-Sequence-like
proof end;
end;

definition
mode T-Sequence is T-Sequence-like Function;
end;

definition
let c1 be set ;
mode T-Sequence of c1 -> T-Sequence means :Def8: :: ORDINAL1:def 8
rng a2 c= a1;
existence
ex b1 being T-Sequence st rng b1 c= c1
proof end;
end;

:: deftheorem Def8 defines T-Sequence ORDINAL1:def 8 :
for b1 being set
for b2 being T-Sequence holds
( b2 is T-Sequence of b1 iff rng b2 c= b1 );

theorem Th43: :: ORDINAL1:43
canceled;

theorem Th44: :: ORDINAL1:44
canceled;

theorem Th45: :: ORDINAL1:45
for b1 being set holds
{} is T-Sequence of b1
proof end;

theorem Th46: :: ORDINAL1:46
for b1 being Function holds
( dom b1 is Ordinal implies b1 is T-Sequence of rng b1 )
proof end;

registration
let c1 be T-Sequence;
cluster dom a1 -> epsilon-transitive epsilon-connected ordinal ;
coherence
dom c1 is ordinal
by Def7;
end;

theorem Th47: :: ORDINAL1:47
for b1, b2 being set holds
( b1 c= b2 implies for b3 being T-Sequence of b1 holds
b3 is T-Sequence of b2 )
proof end;

definition
let c1 be T-Sequence;
let c2 be Ordinal;
redefine func | as c1 | c2 -> T-Sequence of rng a1;
coherence
c1 | c2 is T-Sequence of rng c1
proof end;
end;

theorem Th48: :: ORDINAL1:48
for b1 being set
for b2 being T-Sequence of b1
for b3 being Ordinal holds
b2 | b3 is T-Sequence of b1
proof end;

definition
let c1 be set ;
attr a1 is c=-linear means :Def9: :: ORDINAL1:def 9
for b1, b2 being set holds
( b1 in a1 & b2 in a1 implies b1,b2 are_c=-comparable );
end;

:: deftheorem Def9 defines c=-linear ORDINAL1:def 9 :
for b1 being set holds
( b1 is c=-linear iff for b2, b3 being set holds
( b2 in b1 & b3 in b1 implies b2,b3 are_c=-comparable ) );

theorem Th49: :: ORDINAL1:49
for b1 being set holds
( ( for b2 being set holds
( b2 in b1 implies b2 is T-Sequence ) ) & b1 is c=-linear implies union b1 is T-Sequence )
proof end;

scheme :: ORDINAL1:sch 3
s3{ F1() -> Ordinal, F2( T-Sequence) -> set , F3() -> T-Sequence, F4() -> T-Sequence } :
F3() = F4()
provided
E32: ( dom F3() = F1() & ( for b1 being Ordinal
for b2 being T-Sequence holds
( b1 in F1() & b2 = F3() | b1 implies F3() . b1 = F2(b2) ) ) ) and E33: ( dom F4() = F1() & ( for b1 being Ordinal
for b2 being T-Sequence holds
( b1 in F1() & b2 = F4() | b1 implies F4() . b1 = F2(b2) ) ) )
proof end;

scheme :: ORDINAL1:sch 4
s4{ F1() -> Ordinal, F2( T-Sequence) -> set } :
ex b1 being T-Sequence st
( dom b1 = F1() & ( for b2 being Ordinal
for b3 being T-Sequence holds
( b2 in F1() & b3 = b1 | b2 implies b1 . b2 = F2(b3) ) ) )
proof end;

scheme :: ORDINAL1:sch 5
s5{ F1() -> T-Sequence, F2( Ordinal) -> set , F3( T-Sequence) -> set } :
for b1 being Ordinal holds
( b1 in dom F1() implies F1() . b1 = F3((F1() | b1)) )
provided
E32: for b1 being Ordinal
for b2 being set holds
( b2 = F2(b1) iff ex b3 being T-Sequence st
( b2 = F3(b3) & dom b3 = b1 & ( for b4 being Ordinal holds
( b4 in b1 implies b3 . b4 = F3((b3 | b4)) ) ) ) ) and E33: for b1 being Ordinal holds
( b1 in dom F1() implies F1() . b1 = F2(b1) )
proof end;

theorem Th50: :: ORDINAL1:50
for b1, b2 being Ordinal holds
not ( not b1 c< b2 & not b1 = b2 & not b2 c< b1 )
proof end;

definition
let c1 be set ;
func On c1 -> set means :Def10: :: ORDINAL1:def 10
for b1 being set holds
( b1 in a2 iff ( b1 in a1 & b1 is Ordinal ) );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff ( b2 in c1 & b2 is Ordinal ) )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff ( b3 in c1 & b3 is Ordinal ) ) ) & ( for b3 being set holds
( b3 in b2 iff ( b3 in c1 & b3 is Ordinal ) ) ) implies b1 = b2 )
proof end;
func Lim c1 -> set means :: ORDINAL1:def 11
for b1 being set holds
( b1 in a2 iff ( b1 in a1 & ex b2 being Ordinal st
( b1 = b2 & b2 is_limit_ordinal ) ) );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff ( b2 in c1 & ex b3 being Ordinal st
( b2 = b3 & b3 is_limit_ordinal ) ) )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff ( b3 in c1 & ex b4 being Ordinal st
( b3 = b4 & b4 is_limit_ordinal ) ) ) ) & ( for b3 being set holds
( b3 in b2 iff ( b3 in c1 & ex b4 being Ordinal st
( b3 = b4 & b4 is_limit_ordinal ) ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def10 defines On ORDINAL1:def 10 :
for b1, b2 being set holds
( b2 = On b1 iff for b3 being set holds
( b3 in b2 iff ( b3 in b1 & b3 is Ordinal ) ) );

:: deftheorem Def11 defines Lim ORDINAL1:def 11 :
for b1, b2 being set holds
( b2 = Lim b1 iff for b3 being set holds
( b3 in b2 iff ( b3 in b1 & ex b4 being Ordinal st
( b3 = b4 & b4 is_limit_ordinal ) ) ) );

theorem Th51: :: ORDINAL1:51
for b1 being Ordinal holds
ex b2 being Ordinal st
( b1 in b2 & b2 is_limit_ordinal )
proof end;

definition
func omega -> set means :Def12: :: ORDINAL1:def 12
( {} in a1 & a1 is_limit_ordinal & a1 is ordinal & ( for b1 being Ordinal holds
( {} in b1 & b1 is_limit_ordinal implies a1 c= b1 ) ) );
existence
ex b1 being set st
( {} in b1 & b1 is_limit_ordinal & b1 is ordinal & ( for b2 being Ordinal holds
( {} in b2 & b2 is_limit_ordinal implies b1 c= b2 ) ) )
proof end;
uniqueness
for b1, b2 being set holds
( {} in b1 & b1 is_limit_ordinal & b1 is ordinal & ( for b3 being Ordinal holds
( {} in b3 & b3 is_limit_ordinal implies b1 c= b3 ) ) & {} in b2 & b2 is_limit_ordinal & b2 is ordinal & ( for b3 being Ordinal holds
( {} in b3 & b3 is_limit_ordinal implies b2 c= b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def12 defines omega ORDINAL1:def 12 :
for b1 being set holds
( b1 = omega iff ( {} in b1 & b1 is_limit_ordinal & b1 is ordinal & ( for b2 being Ordinal holds
( {} in b2 & b2 is_limit_ordinal implies b1 c= b2 ) ) ) );

registration
cluster omega -> non empty epsilon-transitive epsilon-connected ordinal ;
coherence
( not omega is empty & omega is ordinal )
by Def12;
end;

definition
let c1 be set ;
attr a1 is natural means :: ORDINAL1:def 13
a1 in omega ;
end;

:: deftheorem Def13 defines natural ORDINAL1:def 13 :
for b1 being set holds
( b1 is natural iff b1 in omega );

registration
let c1 be Ordinal;
cluster -> epsilon-transitive epsilon-connected ordinal Element of a1;
coherence
for b1 being Element of c1 holds b1 is ordinal
proof end;
end;