:: DICKSON semantic presentation begin theorem Th1: :: DICKSON:1 for g being Function for x being set st dom g = {x} holds g = x .--> (g . x) proof let g be Function; ::_thesis: for x being set st dom g = {x} holds g = x .--> (g . x) let x be set ; ::_thesis: ( dom g = {x} implies g = x .--> (g . x) ) assume A1: dom g = {x} ; ::_thesis: g = x .--> (g . x) now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_g_implies_[a,b]_in_x_.-->_(g_._x)_)_&_(_[a,b]_in_x_.-->_(g_._x)_implies_[a,b]_in_g_)_) let a, b be set ; ::_thesis: ( ( [a,b] in g implies [a,b] in x .--> (g . x) ) & ( [a,b] in x .--> (g . x) implies [a,b] in g ) ) A2: dom (x .--> (g . x)) = {x} by FUNCOP_1:13; hereby ::_thesis: ( [a,b] in x .--> (g . x) implies [a,b] in g ) assume A3: [a,b] in g ; ::_thesis: [a,b] in x .--> (g . x) then A4: a in {x} by A1, FUNCT_1:1; then A5: a = x by TARSKI:def_1; then b = g . x by A3, FUNCT_1:1; then (x .--> (g . x)) . a = b by A5, FUNCOP_1:72; hence [a,b] in x .--> (g . x) by A2, A4, FUNCT_1:1; ::_thesis: verum end; assume A6: [a,b] in x .--> (g . x) ; ::_thesis: [a,b] in g then A7: a in {x} by A2, FUNCT_1:1; then A8: a = x by TARSKI:def_1; b = (x .--> (g . x)) . a by A6, FUNCT_1:1 .= g . a by A8, FUNCOP_1:72 ; hence [a,b] in g by A1, A7, FUNCT_1:1; ::_thesis: verum end; hence g = x .--> (g . x) by RELAT_1:def_2; ::_thesis: verum end; theorem Th2: :: DICKSON:2 for n being Nat holds n c= n + 1 proof let n be Nat; ::_thesis: n c= n + 1 n + 1 = n \/ {n} by AFINSQ_1:2; hence n c= n + 1 by XBOOLE_1:7; ::_thesis: verum end; scheme :: DICKSON:sch 1 FinSegRng2{ F1() -> Element of NAT , F2() -> Element of NAT , F3( set ) -> set , P1[ set ] } : { F3(i) where i is Element of NAT : ( F1() < i & i <= F2() & P1[i] ) } is finite proof set F1 = { F3(i) where i is Element of NAT : ( F1() <= i & i <= F2() & P1[i] ) } ; set F2 = { F3(i) where i is Element of NAT : ( F1() < i & i <= F2() & P1[i] ) } ; A1: { F3(i) where i is Element of NAT : ( F1() <= i & i <= F2() & P1[i] ) } is finite from FINSEQ_1:sch_6(); { F3(i) where i is Element of NAT : ( F1() < i & i <= F2() & P1[i] ) } c= { F3(i) where i is Element of NAT : ( F1() <= i & i <= F2() & P1[i] ) } proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { F3(i) where i is Element of NAT : ( F1() < i & i <= F2() & P1[i] ) } or x in { F3(i) where i is Element of NAT : ( F1() <= i & i <= F2() & P1[i] ) } ) assume x in { F3(i) where i is Element of NAT : ( F1() < i & i <= F2() & P1[i] ) } ; ::_thesis: x in { F3(i) where i is Element of NAT : ( F1() <= i & i <= F2() & P1[i] ) } then ex i being Element of NAT st ( F3(i) = x & F1() < i & i <= F2() & P1[i] ) ; hence x in { F3(i) where i is Element of NAT : ( F1() <= i & i <= F2() & P1[i] ) } ; ::_thesis: verum end; hence { F3(i) where i is Element of NAT : ( F1() < i & i <= F2() & P1[i] ) } is finite by A1; ::_thesis: verum end; theorem Th3: :: DICKSON:3 for X being infinite set ex f being Function of NAT,X st f is one-to-one proof let X be infinite set ; ::_thesis: ex f being Function of NAT,X st f is one-to-one card NAT c= card X by CARD_1:47, CARD_3:85; then consider f being Function such that A1: f is one-to-one and A2: dom f = NAT and A3: rng f c= X by CARD_1:10; now__::_thesis:_for_x_being_set_st_x_in_NAT_holds_ f_._x_in_X let x be set ; ::_thesis: ( x in NAT implies f . x in X ) assume x in NAT ; ::_thesis: f . x in X then f . x in rng f by A2, FUNCT_1:3; hence f . x in X by A3; ::_thesis: verum end; then reconsider f = f as Function of NAT,X by A2, FUNCT_2:3; take f ; ::_thesis: f is one-to-one thus f is one-to-one by A1; ::_thesis: verum end; definition let R be RelStr ; let f be sequence of R; attrf is ascending means :: DICKSON:def 1 for n being Element of NAT holds ( f . (n + 1) <> f . n & [(f . n),(f . (n + 1))] in the InternalRel of R ); end; :: deftheorem defines ascending DICKSON:def_1_:_ for R being RelStr for f being sequence of R holds ( f is ascending iff for n being Element of NAT holds ( f . (n + 1) <> f . n & [(f . n),(f . (n + 1))] in the InternalRel of R ) ); definition let R be RelStr ; let f be sequence of R; attrf is weakly-ascending means :Def2: :: DICKSON:def 2 for n being Element of NAT holds [(f . n),(f . (n + 1))] in the InternalRel of R; end; :: deftheorem Def2 defines weakly-ascending DICKSON:def_2_:_ for R being RelStr for f being sequence of R holds ( f is weakly-ascending iff for n being Element of NAT holds [(f . n),(f . (n + 1))] in the InternalRel of R ); theorem Th4: :: DICKSON:4 for R being non empty transitive RelStr for f being sequence of R st f is weakly-ascending holds for i, j being Element of NAT st i < j holds f . i <= f . j proof let R be non empty transitive RelStr ; ::_thesis: for f being sequence of R st f is weakly-ascending holds for i, j being Element of NAT st i < j holds f . i <= f . j let f be sequence of R; ::_thesis: ( f is weakly-ascending implies for i, j being Element of NAT st i < j holds f . i <= f . j ) assume A1: f is weakly-ascending ; ::_thesis: for i, j being Element of NAT st i < j holds f . i <= f . j let i be Element of NAT ; ::_thesis: for j being Element of NAT st i < j holds f . i <= f . j defpred S1[ Element of NAT ] means ( i < $1 implies f . i <= f . $1 ); A2: S1[ 0 ] by NAT_1:2; A3: for j being Element of NAT st S1[j] holds S1[j + 1] proof let j be Element of NAT ; ::_thesis: ( S1[j] implies S1[j + 1] ) assume that A4: S1[j] and A5: i < j + 1 ; ::_thesis: f . i <= f . (j + 1) reconsider fj1 = f . (j + 1) as Element of R ; A6: [(f . j),(f . (j + 1))] in the InternalRel of R by A1, Def2; then A7: f . j <= fj1 by ORDERS_2:def_5; A8: i <= j by A5, NAT_1:13; percases ( i < j or i = j ) by A8, XXREAL_0:1; suppose i < j ; ::_thesis: f . i <= f . (j + 1) hence f . i <= f . (j + 1) by A4, A7, ORDERS_2:3; ::_thesis: verum end; suppose i = j ; ::_thesis: f . i <= f . (j + 1) hence f . i <= f . (j + 1) by A6, ORDERS_2:def_5; ::_thesis: verum end; end; end; thus for j being Element of NAT holds S1[j] from NAT_1:sch_1(A2, A3); ::_thesis: verum end; theorem Th5: :: DICKSON:5 for R being non empty RelStr holds ( R is connected iff the InternalRel of R is_strongly_connected_in the carrier of R ) proof let R be non empty RelStr ; ::_thesis: ( R is connected iff the InternalRel of R is_strongly_connected_in the carrier of R ) set IR = the InternalRel of R; set CR = the carrier of R; hereby ::_thesis: ( the InternalRel of R is_strongly_connected_in the carrier of R implies R is connected ) assume A1: R is connected ; ::_thesis: the InternalRel of R is_strongly_connected_in the carrier of R now__::_thesis:_for_x,_y_being_set_st_x_in_the_carrier_of_R_&_y_in_the_carrier_of_R_&_not_[x,y]_in_the_InternalRel_of_R_holds_ [y,x]_in_the_InternalRel_of_R let x, y be set ; ::_thesis: ( x in the carrier of R & y in the carrier of R & not [x,y] in the InternalRel of R implies [y,x] in the InternalRel of R ) assume that A2: x in the carrier of R and A3: y in the carrier of R ; ::_thesis: ( [x,y] in the InternalRel of R or [y,x] in the InternalRel of R ) reconsider x9 = x, y9 = y as Element of R by A2, A3; ( x9 <= y9 or y9 <= x9 ) by A1, WAYBEL_0:def_29; hence ( [x,y] in the InternalRel of R or [y,x] in the InternalRel of R ) by ORDERS_2:def_5; ::_thesis: verum end; hence the InternalRel of R is_strongly_connected_in the carrier of R by RELAT_2:def_7; ::_thesis: verum end; assume A4: the InternalRel of R is_strongly_connected_in the carrier of R ; ::_thesis: R is connected now__::_thesis:_for_x,_y_being_Element_of_R_holds_ (_x_<=_y_or_y_<=_x_) let x, y be Element of R; ::_thesis: ( x <= y or y <= x ) ( [x,y] in the InternalRel of R or [y,x] in the InternalRel of R ) by A4, RELAT_2:def_7; hence ( x <= y or y <= x ) by ORDERS_2:def_5; ::_thesis: verum end; hence R is connected by WAYBEL_0:def_29; ::_thesis: verum end; theorem Th6: :: DICKSON:6 for L being RelStr for Y, a being set holds ( ( the InternalRel of L -Seg a misses Y & a in Y ) iff a is_minimal_wrt Y, the InternalRel of L ) proof let L be RelStr ; ::_thesis: for Y, a being set holds ( ( the InternalRel of L -Seg a misses Y & a in Y ) iff a is_minimal_wrt Y, the InternalRel of L ) let Y, a be set ; ::_thesis: ( ( the InternalRel of L -Seg a misses Y & a in Y ) iff a is_minimal_wrt Y, the InternalRel of L ) set IR = the InternalRel of L; hereby ::_thesis: ( a is_minimal_wrt Y, the InternalRel of L implies ( the InternalRel of L -Seg a misses Y & a in Y ) ) assume that A1: the InternalRel of L -Seg a misses Y and A2: a in Y ; ::_thesis: a is_minimal_wrt Y, the InternalRel of L A3: ( the InternalRel of L -Seg a) /\ Y = {} by A1, XBOOLE_0:def_7; now__::_thesis:_for_y_being_set_holds_ (_not_y_in_Y_or_not_y_<>_a_or_not_[y,a]_in_the_InternalRel_of_L_) assume ex y being set st ( y in Y & y <> a & [y,a] in the InternalRel of L ) ; ::_thesis: contradiction then consider y being set such that A4: y in Y and A5: y <> a and A6: [y,a] in the InternalRel of L ; y in the InternalRel of L -Seg a by A5, A6, WELLORD1:1; hence contradiction by A3, A4, XBOOLE_0:def_4; ::_thesis: verum end; hence a is_minimal_wrt Y, the InternalRel of L by A2, WAYBEL_4:def_25; ::_thesis: verum end; assume A7: a is_minimal_wrt Y, the InternalRel of L ; ::_thesis: ( the InternalRel of L -Seg a misses Y & a in Y ) now__::_thesis:_the_InternalRel_of_L_-Seg_a_misses_Y assume not the InternalRel of L -Seg a misses Y ; ::_thesis: contradiction then ( the InternalRel of L -Seg a) /\ Y <> {} by XBOOLE_0:def_7; then consider y being set such that A8: y in ( the InternalRel of L -Seg a) /\ Y by XBOOLE_0:def_1; A9: y in the InternalRel of L -Seg a by A8, XBOOLE_0:def_4; A10: y in Y by A8, XBOOLE_0:def_4; A11: y <> a by A9, WELLORD1:1; [y,a] in the InternalRel of L by A9, WELLORD1:1; hence contradiction by A7, A10, A11, WAYBEL_4:def_25; ::_thesis: verum end; hence the InternalRel of L -Seg a misses Y ; ::_thesis: a in Y thus a in Y by A7, WAYBEL_4:def_25; ::_thesis: verum end; theorem Th7: :: DICKSON:7 for L being non empty transitive antisymmetric RelStr for x being Element of L for a, N being set st a is_minimal_wrt ( the InternalRel of L -Seg x) /\ N, the InternalRel of L holds a is_minimal_wrt N, the InternalRel of L proof let L be non empty transitive antisymmetric RelStr ; ::_thesis: for x being Element of L for a, N being set st a is_minimal_wrt ( the InternalRel of L -Seg x) /\ N, the InternalRel of L holds a is_minimal_wrt N, the InternalRel of L let x be Element of L; ::_thesis: for a, N being set st a is_minimal_wrt ( the InternalRel of L -Seg x) /\ N, the InternalRel of L holds a is_minimal_wrt N, the InternalRel of L let a, N be set ; ::_thesis: ( a is_minimal_wrt ( the InternalRel of L -Seg x) /\ N, the InternalRel of L implies a is_minimal_wrt N, the InternalRel of L ) assume A1: a is_minimal_wrt ( the InternalRel of L -Seg x) /\ N, the InternalRel of L ; ::_thesis: a is_minimal_wrt N, the InternalRel of L set IR = the InternalRel of L; set CR = the carrier of L; A2: the InternalRel of L is_transitive_in the carrier of L by ORDERS_2:def_3; now__::_thesis:_(_a_in_N_&_(_for_y_being_set_holds_ (_not_y_in_N_or_not_y_<>_a_or_not_[y,a]_in_the_InternalRel_of_L_)_)_) A3: a in ( the InternalRel of L -Seg x) /\ N by A1, WAYBEL_4:def_25; then A4: a in the InternalRel of L -Seg x by XBOOLE_0:def_4; then A5: a <> x by WELLORD1:1; A6: [a,x] in the InternalRel of L by A4, WELLORD1:1; then reconsider a9 = a as Element of L by ZFMISC_1:87; thus a in N by A3, XBOOLE_0:def_4; ::_thesis: for y being set holds ( not y in N or not y <> a or not [y,a] in the InternalRel of L ) now__::_thesis:_for_y_being_set_holds_ (_not_y_in_N_or_not_y_<>_a_or_not_[y,a]_in_the_InternalRel_of_L_) assume ex y being set st ( y in N & y <> a & [y,a] in the InternalRel of L ) ; ::_thesis: contradiction then consider y being set such that A7: y in N and A8: y <> a and A9: [y,a] in the InternalRel of L ; A10: a in the carrier of L by A9, ZFMISC_1:87; y in the carrier of L by A9, ZFMISC_1:87; then A11: [y,x] in the InternalRel of L by A2, A6, A9, A10, RELAT_2:def_8; percases ( x = y or x <> y ) ; suppose x = y ; ::_thesis: contradiction then A12: x <= a9 by A9, ORDERS_2:def_5; a9 <= x by A6, ORDERS_2:def_5; hence contradiction by A5, A12, ORDERS_2:2; ::_thesis: verum end; suppose x <> y ; ::_thesis: contradiction then y in the InternalRel of L -Seg x by A11, WELLORD1:1; then y in ( the InternalRel of L -Seg x) /\ N by A7, XBOOLE_0:def_4; hence contradiction by A1, A8, A9, WAYBEL_4:def_25; ::_thesis: verum end; end; end; hence for y being set holds ( not y in N or not y <> a or not [y,a] in the InternalRel of L ) ; ::_thesis: verum end; hence a is_minimal_wrt N, the InternalRel of L by WAYBEL_4:def_25; ::_thesis: verum end; begin definition let R be RelStr ; attrR is quasi_ordered means :Def3: :: DICKSON:def 3 ( R is reflexive & R is transitive ); end; :: deftheorem Def3 defines quasi_ordered DICKSON:def_3_:_ for R being RelStr holds ( R is quasi_ordered iff ( R is reflexive & R is transitive ) ); definition let R be RelStr ; assume B1: R is quasi_ordered ; func EqRel R -> Equivalence_Relation of the carrier of R equals :Def4: :: DICKSON:def 4 the InternalRel of R /\ ( the InternalRel of R ~); coherence the InternalRel of R /\ ( the InternalRel of R ~) is Equivalence_Relation of the carrier of R proof set IR = the InternalRel of R; set CR = the carrier of R; R is reflexive by B1, Def3; then A1: the InternalRel of R is_reflexive_in the carrier of R by ORDERS_2:def_2; R is transitive by B1, Def3; then A2: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; then A3: the InternalRel of R quasi_orders the carrier of R by A1, ORDERS_1:def_6; A4: the InternalRel of R /\ ( the InternalRel of R ~) is_reflexive_in the carrier of R proof let x be set ; :: according to RELAT_2:def_1 ::_thesis: ( not x in the carrier of R or [x,x] in the InternalRel of R /\ ( the InternalRel of R ~) ) assume x in the carrier of R ; ::_thesis: [x,x] in the InternalRel of R /\ ( the InternalRel of R ~) then A5: [x,x] in the InternalRel of R by A1, RELAT_2:def_1; then [x,x] in the InternalRel of R ~ by RELAT_1:def_7; hence [x,x] in the InternalRel of R /\ ( the InternalRel of R ~) by A5, XBOOLE_0:def_4; ::_thesis: verum end; then A6: dom ( the InternalRel of R /\ ( the InternalRel of R ~)) = the carrier of R by ORDERS_1:13; A7: field ( the InternalRel of R /\ ( the InternalRel of R ~)) = the carrier of R by A4, ORDERS_1:13; A8: the InternalRel of R /\ ( the InternalRel of R ~) is_symmetric_in the carrier of R proof let x, y be set ; :: according to RELAT_2:def_3 ::_thesis: ( not x in the carrier of R or not y in the carrier of R or not [x,y] in the InternalRel of R /\ ( the InternalRel of R ~) or [y,x] in the InternalRel of R /\ ( the InternalRel of R ~) ) assume that x in the carrier of R and y in the carrier of R and A9: [x,y] in the InternalRel of R /\ ( the InternalRel of R ~) ; ::_thesis: [y,x] in the InternalRel of R /\ ( the InternalRel of R ~) A10: [x,y] in the InternalRel of R by A9, XBOOLE_0:def_4; A11: [x,y] in the InternalRel of R ~ by A9, XBOOLE_0:def_4; A12: [y,x] in the InternalRel of R ~ by A10, RELAT_1:def_7; [y,x] in the InternalRel of R by A11, RELAT_1:def_7; hence [y,x] in the InternalRel of R /\ ( the InternalRel of R ~) by A12, XBOOLE_0:def_4; ::_thesis: verum end; the InternalRel of R /\ ( the InternalRel of R ~) is_transitive_in the carrier of R proof let x, y, z be set ; :: according to RELAT_2:def_8 ::_thesis: ( not x in the carrier of R or not y in the carrier of R or not z in the carrier of R or not [x,y] in the InternalRel of R /\ ( the InternalRel of R ~) or not [y,z] in the InternalRel of R /\ ( the InternalRel of R ~) or [x,z] in the InternalRel of R /\ ( the InternalRel of R ~) ) assume that A13: x in the carrier of R and A14: y in the carrier of R and A15: z in the carrier of R and A16: [x,y] in the InternalRel of R /\ ( the InternalRel of R ~) and A17: [y,z] in the InternalRel of R /\ ( the InternalRel of R ~) ; ::_thesis: [x,z] in the InternalRel of R /\ ( the InternalRel of R ~) A18: [x,y] in the InternalRel of R by A16, XBOOLE_0:def_4; A19: [x,y] in the InternalRel of R ~ by A16, XBOOLE_0:def_4; A20: [y,z] in the InternalRel of R by A17, XBOOLE_0:def_4; A21: [y,z] in the InternalRel of R ~ by A17, XBOOLE_0:def_4; A22: [x,z] in the InternalRel of R by A2, A13, A14, A15, A18, A20, RELAT_2:def_8; the InternalRel of R ~ quasi_orders the carrier of R by A3, ORDERS_1:40; then the InternalRel of R ~ is_transitive_in the carrier of R by ORDERS_1:def_6; then [x,z] in the InternalRel of R ~ by A13, A14, A15, A19, A21, RELAT_2:def_8; hence [x,z] in the InternalRel of R /\ ( the InternalRel of R ~) by A22, XBOOLE_0:def_4; ::_thesis: verum end; hence the InternalRel of R /\ ( the InternalRel of R ~) is Equivalence_Relation of the carrier of R by A6, A7, A8, PARTFUN1:def_2, RELAT_2:def_11, RELAT_2:def_16; ::_thesis: verum end; end; :: deftheorem Def4 defines EqRel DICKSON:def_4_:_ for R being RelStr st R is quasi_ordered holds EqRel R = the InternalRel of R /\ ( the InternalRel of R ~); theorem Th8: :: DICKSON:8 for R being RelStr for x, y being Element of R st R is quasi_ordered holds ( x in Class ((EqRel R),y) iff ( x <= y & y <= x ) ) proof let R be RelStr ; ::_thesis: for x, y being Element of R st R is quasi_ordered holds ( x in Class ((EqRel R),y) iff ( x <= y & y <= x ) ) let x, y be Element of R; ::_thesis: ( R is quasi_ordered implies ( x in Class ((EqRel R),y) iff ( x <= y & y <= x ) ) ) assume A1: R is quasi_ordered ; ::_thesis: ( x in Class ((EqRel R),y) iff ( x <= y & y <= x ) ) set IR = the InternalRel of R; hereby ::_thesis: ( x <= y & y <= x implies x in Class ((EqRel R),y) ) assume x in Class ((EqRel R),y) ; ::_thesis: ( x <= y & y <= x ) then [x,y] in EqRel R by EQREL_1:19; then A2: [x,y] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; then A3: [x,y] in the InternalRel of R by XBOOLE_0:def_4; A4: [x,y] in the InternalRel of R ~ by A2, XBOOLE_0:def_4; thus x <= y by A3, ORDERS_2:def_5; ::_thesis: y <= x [y,x] in the InternalRel of R by A4, RELAT_1:def_7; hence y <= x by ORDERS_2:def_5; ::_thesis: verum end; assume that A5: x <= y and A6: y <= x ; ::_thesis: x in Class ((EqRel R),y) A7: [y,x] in the InternalRel of R by A6, ORDERS_2:def_5; A8: [x,y] in the InternalRel of R by A5, ORDERS_2:def_5; [x,y] in the InternalRel of R ~ by A7, RELAT_1:def_7; then [x,y] in the InternalRel of R /\ ( the InternalRel of R ~) by A8, XBOOLE_0:def_4; then [x,y] in EqRel R by A1, Def4; hence x in Class ((EqRel R),y) by EQREL_1:19; ::_thesis: verum end; definition let R be RelStr ; func <=E R -> Relation of (Class (EqRel R)) means :Def5: :: DICKSON:def 5 for A, B being set holds ( [A,B] in it iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ); existence ex b1 being Relation of (Class (EqRel R)) st for A, B being set holds ( [A,B] in b1 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) proof set IR = the InternalRel of R; set CR = the carrier of R; percases ( the carrier of R = {} or not the carrier of R is empty ) ; supposeA1: the carrier of R = {} ; ::_thesis: ex b1 being Relation of (Class (EqRel R)) st for A, B being set holds ( [A,B] in b1 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) reconsider IT = {} as Relation of (Class (EqRel R)) by RELSET_1:12; take IT ; ::_thesis: for A, B being set holds ( [A,B] in IT iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) let A, B be set ; ::_thesis: ( [A,B] in IT iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) thus ( [A,B] in IT implies ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ; ::_thesis: ( ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) implies [A,B] in IT ) given a, b being Element of R such that A = Class ((EqRel R),a) and B = Class ((EqRel R),b) and A2: a <= b ; ::_thesis: [A,B] in IT the InternalRel of R = {} by A1; hence [A,B] in IT by A2, ORDERS_2:def_5; ::_thesis: verum end; suppose not the carrier of R is empty ; ::_thesis: ex b1 being Relation of (Class (EqRel R)) st for A, B being set holds ( [A,B] in b1 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) then reconsider R9 = R as non empty RelStr ; set IT = { [(Class ((EqRel R),a)),(Class ((EqRel R),b))] where a, b is Element of R9 : a <= b } ; set Y = Class (EqRel R); { [(Class ((EqRel R),a)),(Class ((EqRel R),b))] where a, b is Element of R9 : a <= b } c= [:(Class (EqRel R)),(Class (EqRel R)):] proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { [(Class ((EqRel R),a)),(Class ((EqRel R),b))] where a, b is Element of R9 : a <= b } or x in [:(Class (EqRel R)),(Class (EqRel R)):] ) assume x in { [(Class ((EqRel R),a)),(Class ((EqRel R),b))] where a, b is Element of R9 : a <= b } ; ::_thesis: x in [:(Class (EqRel R)),(Class (EqRel R)):] then consider a, b being Element of R9 such that A3: x = [(Class ((EqRel R),a)),(Class ((EqRel R),b))] and a <= b ; A4: Class ((EqRel R),a) in Class (EqRel R) by EQREL_1:def_3; Class ((EqRel R),b) in Class (EqRel R) by EQREL_1:def_3; hence x in [:(Class (EqRel R)),(Class (EqRel R)):] by A3, A4, ZFMISC_1:def_2; ::_thesis: verum end; then reconsider IT = { [(Class ((EqRel R),a)),(Class ((EqRel R),b))] where a, b is Element of R9 : a <= b } as Relation of (Class (EqRel R)) ; take IT ; ::_thesis: for A, B being set holds ( [A,B] in IT iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) let A, B be set ; ::_thesis: ( [A,B] in IT iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) hereby ::_thesis: ( ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) implies [A,B] in IT ) assume [A,B] in IT ; ::_thesis: ex a9, b9 being Element of R st ( A = Class ((EqRel R),a9) & B = Class ((EqRel R),b9) & a9 <= b9 ) then consider a, b being Element of R such that A5: [A,B] = [(Class ((EqRel R),a)),(Class ((EqRel R),b))] and A6: a <= b ; reconsider a9 = a, b9 = b as Element of R ; take a9 = a9; ::_thesis: ex b9 being Element of R st ( A = Class ((EqRel R),a9) & B = Class ((EqRel R),b9) & a9 <= b9 ) take b9 = b9; ::_thesis: ( A = Class ((EqRel R),a9) & B = Class ((EqRel R),b9) & a9 <= b9 ) thus ( A = Class ((EqRel R),a9) & B = Class ((EqRel R),b9) & a9 <= b9 ) by A5, A6, XTUPLE_0:1; ::_thesis: verum end; given a, b being Element of R such that A7: A = Class ((EqRel R),a) and A8: B = Class ((EqRel R),b) and A9: a <= b ; ::_thesis: [A,B] in IT thus [A,B] in IT by A7, A8, A9; ::_thesis: verum end; end; end; uniqueness for b1, b2 being Relation of (Class (EqRel R)) st ( for A, B being set holds ( [A,B] in b1 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ) & ( for A, B being set holds ( [A,B] in b2 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ) holds b1 = b2 proof let IT1, IT2 be Relation of (Class (EqRel R)); ::_thesis: ( ( for A, B being set holds ( [A,B] in IT1 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ) & ( for A, B being set holds ( [A,B] in IT2 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ) implies IT1 = IT2 ) assume that A10: for A, B being set holds ( [A,B] in IT1 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) and A11: for A, B being set holds ( [A,B] in IT2 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_x_being_set_holds_ (_(_x_in_IT1_implies_x_in_IT2_)_&_(_x_in_IT2_implies_x_in_IT1_)_) let x be set ; ::_thesis: ( ( x in IT1 implies x in IT2 ) & ( x in IT2 implies x in IT1 ) ) hereby ::_thesis: ( x in IT2 implies x in IT1 ) assume A12: x in IT1 ; ::_thesis: x in IT2 set Y = Class (EqRel R); consider A, B being set such that A in Class (EqRel R) and B in Class (EqRel R) and A13: x = [A,B] by A12, ZFMISC_1:def_2; ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) by A10, A12, A13; hence x in IT2 by A11, A13; ::_thesis: verum end; assume A14: x in IT2 ; ::_thesis: x in IT1 set Y = Class (EqRel R); consider A, B being set such that A in Class (EqRel R) and B in Class (EqRel R) and A15: x = [A,B] by A14, ZFMISC_1:def_2; ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) by A11, A14, A15; hence x in IT1 by A10, A15; ::_thesis: verum end; hence IT1 = IT2 by TARSKI:1; ::_thesis: verum end; end; :: deftheorem Def5 defines <=E DICKSON:def_5_:_ for R being RelStr for b2 being Relation of (Class (EqRel R)) holds ( b2 = <=E R iff for A, B being set holds ( [A,B] in b2 iff ex a, b being Element of R st ( A = Class ((EqRel R),a) & B = Class ((EqRel R),b) & a <= b ) ) ); theorem Th9: :: DICKSON:9 for R being RelStr st R is quasi_ordered holds <=E R partially_orders Class (EqRel R) proof let R be RelStr ; ::_thesis: ( R is quasi_ordered implies <=E R partially_orders Class (EqRel R) ) set CR = the carrier of R; set IR = the InternalRel of R; assume A1: R is quasi_ordered ; ::_thesis: <=E R partially_orders Class (EqRel R) then R is transitive by Def3; then A2: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; thus <=E R is_reflexive_in Class (EqRel R) :: according to ORDERS_1:def_7 ::_thesis: ( <=E R is_transitive_in Class (EqRel R) & <=E R is_antisymmetric_in Class (EqRel R) ) proof let x be set ; :: according to RELAT_2:def_1 ::_thesis: ( not x in Class (EqRel R) or [x,x] in <=E R ) assume x in Class (EqRel R) ; ::_thesis: [x,x] in <=E R then consider a being set such that A3: a in the carrier of R and A4: x = Class ((EqRel R),a) by EQREL_1:def_3; R is reflexive by A1, Def3; then the InternalRel of R is_reflexive_in the carrier of R by ORDERS_2:def_2; then A5: [a,a] in the InternalRel of R by A3, RELAT_2:def_1; reconsider a9 = a as Element of R by A3; a9 <= a9 by A5, ORDERS_2:def_5; hence [x,x] in <=E R by A4, Def5; ::_thesis: verum end; thus <=E R is_transitive_in Class (EqRel R) ::_thesis: <=E R is_antisymmetric_in Class (EqRel R) proof let x, y, z be set ; :: according to RELAT_2:def_8 ::_thesis: ( not x in Class (EqRel R) or not y in Class (EqRel R) or not z in Class (EqRel R) or not [x,y] in <=E R or not [y,z] in <=E R or [x,z] in <=E R ) assume that A6: x in Class (EqRel R) and y in Class (EqRel R) and z in Class (EqRel R) and A7: [x,y] in <=E R and A8: [y,z] in <=E R ; ::_thesis: [x,z] in <=E R consider a, b being Element of R such that A9: x = Class ((EqRel R),a) and A10: y = Class ((EqRel R),b) and A11: a <= b by A7, Def5; consider c, d being Element of R such that A12: y = Class ((EqRel R),c) and A13: z = Class ((EqRel R),d) and A14: c <= d by A8, Def5; A15: [a,b] in the InternalRel of R by A11, ORDERS_2:def_5; A16: [c,d] in the InternalRel of R by A14, ORDERS_2:def_5; A17: ex x1 being set st ( x1 in the carrier of R & x = Class ((EqRel R),x1) ) by A6, EQREL_1:def_3; then b in Class ((EqRel R),c) by A10, A12, EQREL_1:23; then [b,c] in EqRel R by EQREL_1:19; then [b,c] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; then [b,c] in the InternalRel of R by XBOOLE_0:def_4; then [a,c] in the InternalRel of R by A2, A15, A17, RELAT_2:def_8; then [a,d] in the InternalRel of R by A2, A16, A17, RELAT_2:def_8; then a <= d by ORDERS_2:def_5; hence [x,z] in <=E R by A9, A13, Def5; ::_thesis: verum end; thus <=E R is_antisymmetric_in Class (EqRel R) ::_thesis: verum proof let x, y be set ; :: according to RELAT_2:def_4 ::_thesis: ( not x in Class (EqRel R) or not y in Class (EqRel R) or not [x,y] in <=E R or not [y,x] in <=E R or x = y ) assume that A18: x in Class (EqRel R) and y in Class (EqRel R) and A19: [x,y] in <=E R and A20: [y,x] in <=E R ; ::_thesis: x = y consider a, b being Element of R such that A21: x = Class ((EqRel R),a) and A22: y = Class ((EqRel R),b) and A23: a <= b by A19, Def5; consider c, d being Element of R such that A24: y = Class ((EqRel R),c) and A25: x = Class ((EqRel R),d) and A26: c <= d by A20, Def5; A27: [a,b] in the InternalRel of R by A23, ORDERS_2:def_5; A28: [c,d] in the InternalRel of R by A26, ORDERS_2:def_5; A29: ex x1 being set st ( x1 in the carrier of R & x = Class ((EqRel R),x1) ) by A18, EQREL_1:def_3; then A30: d in Class ((EqRel R),a) by A21, A25, EQREL_1:23; a in Class ((EqRel R),a) by A29, EQREL_1:20; then A31: [a,d] in EqRel R by A30, EQREL_1:22; A32: c in Class ((EqRel R),b) by A22, A24, A29, EQREL_1:23; b in Class ((EqRel R),b) by A29, EQREL_1:20; then A33: [b,c] in EqRel R by A32, EQREL_1:22; [a,d] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, A31, Def4; then [a,d] in the InternalRel of R ~ by XBOOLE_0:def_4; then A34: [d,a] in the InternalRel of R by RELAT_1:def_7; [b,c] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, A33, Def4; then [b,c] in the InternalRel of R by XBOOLE_0:def_4; then [b,d] in the InternalRel of R by A2, A28, A29, RELAT_2:def_8; then A35: [b,a] in the InternalRel of R by A2, A29, A34, RELAT_2:def_8; [b,a] in the InternalRel of R ~ by A27, RELAT_1:def_7; then [b,a] in the InternalRel of R /\ ( the InternalRel of R ~) by A35, XBOOLE_0:def_4; then [b,a] in EqRel R by A1, Def4; then b in Class ((EqRel R),a) by EQREL_1:19; hence x = y by A21, A22, EQREL_1:23; ::_thesis: verum end; end; theorem :: DICKSON:10 for R being non empty RelStr st R is quasi_ordered & R is connected holds <=E R linearly_orders Class (EqRel R) proof let R be non empty RelStr ; ::_thesis: ( R is quasi_ordered & R is connected implies <=E R linearly_orders Class (EqRel R) ) assume that A1: R is quasi_ordered and A2: R is connected ; ::_thesis: <=E R linearly_orders Class (EqRel R) A3: <=E R partially_orders Class (EqRel R) by A1, Th9; hence <=E R is_reflexive_in Class (EqRel R) by ORDERS_1:def_7; :: according to ORDERS_1:def_8 ::_thesis: ( <=E R is_transitive_in Class (EqRel R) & <=E R is_antisymmetric_in Class (EqRel R) & <=E R is_connected_in Class (EqRel R) ) thus <=E R is_transitive_in Class (EqRel R) by A3, ORDERS_1:def_7; ::_thesis: ( <=E R is_antisymmetric_in Class (EqRel R) & <=E R is_connected_in Class (EqRel R) ) thus <=E R is_antisymmetric_in Class (EqRel R) by A3, ORDERS_1:def_7; ::_thesis: <=E R is_connected_in Class (EqRel R) thus <=E R is_connected_in Class (EqRel R) ::_thesis: verum proof set CR = the carrier of R; let x, y be set ; :: according to RELAT_2:def_6 ::_thesis: ( not x in Class (EqRel R) or not y in Class (EqRel R) or x = y or [x,y] in <=E R or [y,x] in <=E R ) assume that A4: x in Class (EqRel R) and A5: y in Class (EqRel R) and x <> y and A6: not [x,y] in <=E R ; ::_thesis: [y,x] in <=E R consider a being set such that A7: a in the carrier of R and A8: x = Class ((EqRel R),a) by A4, EQREL_1:def_3; consider b being set such that A9: b in the carrier of R and A10: y = Class ((EqRel R),b) by A5, EQREL_1:def_3; reconsider a9 = a, b9 = b as Element of R by A7, A9; not a9 <= b9 by A6, A8, A10, Def5; then b9 <= a9 by A2, WAYBEL_0:def_29; hence [y,x] in <=E R by A8, A10, Def5; ::_thesis: verum end; end; definition let R be Relation; funcR \~ -> Relation equals :: DICKSON:def 6 R \ (R ~); correctness coherence R \ (R ~) is Relation; ; end; :: deftheorem defines \~ DICKSON:def_6_:_ for R being Relation holds R \~ = R \ (R ~); registration let R be Relation; clusterR \~ -> asymmetric ; coherence R \~ is asymmetric proof now__::_thesis:_for_x,_y_being_set_st_x_in_field_(R_\~)_&_y_in_field_(R_\~)_&_[x,y]_in_R_\~_holds_ not_[y,x]_in_R_\~ let x, y be set ; ::_thesis: ( x in field (R \~) & y in field (R \~) & [x,y] in R \~ implies not [y,x] in R \~ ) assume that x in field (R \~) and y in field (R \~) and A1: [x,y] in R \~ ; ::_thesis: not [y,x] in R \~ not [x,y] in R ~ by A1, XBOOLE_0:def_5; hence not [y,x] in R \~ by RELAT_1:def_7; ::_thesis: verum end; then R \~ is_asymmetric_in field (R \~) by RELAT_2:def_5; hence R \~ is asymmetric by RELAT_2:def_13; ::_thesis: verum end; end; definition let X be set ; let R be Relation of X; :: original: \~ redefine funcR \~ -> Relation of X; coherence R \~ is Relation of X proof R \~ = R \ (R ~) ; hence R \~ is Relation of X ; ::_thesis: verum end; end; definition let R be RelStr ; funcR \~ -> strict RelStr equals :: DICKSON:def 7 RelStr(# the carrier of R,( the InternalRel of R \~) #); correctness coherence RelStr(# the carrier of R,( the InternalRel of R \~) #) is strict RelStr ; ; end; :: deftheorem defines \~ DICKSON:def_7_:_ for R being RelStr holds R \~ = RelStr(# the carrier of R,( the InternalRel of R \~) #); registration let R be non empty RelStr ; clusterR \~ -> non empty strict ; coherence not R \~ is empty ; end; registration let R be transitive RelStr ; clusterR \~ -> strict transitive ; coherence R \~ is transitive proof set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); set CR9 = the carrier of (R \~); A1: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; now__::_thesis:_for_x,_y,_z_being_set_st_x_in_the_carrier_of_(R_\~)_&_y_in_the_carrier_of_(R_\~)_&_z_in_the_carrier_of_(R_\~)_&_[x,y]_in_the_InternalRel_of_(R_\~)_&_[y,z]_in_the_InternalRel_of_(R_\~)_holds_ [x,z]_in_the_InternalRel_of_(R_\~) let x, y, z be set ; ::_thesis: ( x in the carrier of (R \~) & y in the carrier of (R \~) & z in the carrier of (R \~) & [x,y] in the InternalRel of (R \~) & [y,z] in the InternalRel of (R \~) implies [x,z] in the InternalRel of (R \~) ) assume that A2: x in the carrier of (R \~) and A3: y in the carrier of (R \~) and A4: z in the carrier of (R \~) and A5: [x,y] in the InternalRel of (R \~) and A6: [y,z] in the InternalRel of (R \~) ; ::_thesis: [x,z] in the InternalRel of (R \~) A7: not [x,y] in the InternalRel of R ~ by A5, XBOOLE_0:def_5; A8: [x,z] in the InternalRel of R by A1, A2, A3, A4, A5, A6, RELAT_2:def_8; now__::_thesis:_not_[x,z]_in_the_InternalRel_of_R_~ assume [x,z] in the InternalRel of R ~ ; ::_thesis: contradiction then [z,x] in the InternalRel of R by RELAT_1:def_7; then [y,x] in the InternalRel of R by A1, A2, A3, A4, A6, RELAT_2:def_8; hence contradiction by A7, RELAT_1:def_7; ::_thesis: verum end; hence [x,z] in the InternalRel of (R \~) by A8, XBOOLE_0:def_5; ::_thesis: verum end; then the InternalRel of (R \~) is_transitive_in the carrier of (R \~) by RELAT_2:def_8; hence R \~ is transitive by ORDERS_2:def_3; ::_thesis: verum end; end; registration let R be RelStr ; clusterR \~ -> strict antisymmetric ; coherence R \~ is antisymmetric proof set IR = the InternalRel of R; set IR9 = the InternalRel of (R \~); set CR9 = the carrier of (R \~); now__::_thesis:_for_x,_y_being_set_st_x_in_the_carrier_of_(R_\~)_&_y_in_the_carrier_of_(R_\~)_&_[x,y]_in_the_InternalRel_of_(R_\~)_&_[y,x]_in_the_InternalRel_of_(R_\~)_holds_ x_=_y let x, y be set ; ::_thesis: ( x in the carrier of (R \~) & y in the carrier of (R \~) & [x,y] in the InternalRel of (R \~) & [y,x] in the InternalRel of (R \~) implies x = y ) assume that x in the carrier of (R \~) and y in the carrier of (R \~) and A1: [x,y] in the InternalRel of (R \~) and A2: [y,x] in the InternalRel of (R \~) ; ::_thesis: x = y not [x,y] in the InternalRel of R ~ by A1, XBOOLE_0:def_5; hence x = y by A2, RELAT_1:def_7; ::_thesis: verum end; then the InternalRel of (R \~) is_antisymmetric_in the carrier of (R \~) by RELAT_2:def_4; hence R \~ is antisymmetric by ORDERS_2:def_4; ::_thesis: verum end; end; theorem :: DICKSON:11 for R being non empty Poset for x being Element of R holds Class ((EqRel R),x) = {x} proof let R be non empty Poset; ::_thesis: for x being Element of R holds Class ((EqRel R),x) = {x} set IR = the InternalRel of R; set CR = the carrier of R; let x be Element of the carrier of R; ::_thesis: Class ((EqRel R),x) = {x} A1: R is quasi_ordered by Def3; A2: the InternalRel of R is_antisymmetric_in the carrier of R by ORDERS_2:def_4; now__::_thesis:_for_z_being_set_holds_ (_(_z_in_Class_((EqRel_R),x)_implies_z_in_{x}_)_&_(_z_in_{x}_implies_z_in_Class_((EqRel_R),x)_)_) let z be set ; ::_thesis: ( ( z in Class ((EqRel R),x) implies z in {x} ) & ( z in {x} implies z in Class ((EqRel R),x) ) ) hereby ::_thesis: ( z in {x} implies z in Class ((EqRel R),x) ) assume z in Class ((EqRel R),x) ; ::_thesis: z in {x} then [z,x] in EqRel R by EQREL_1:19; then A3: [z,x] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; then A4: [z,x] in the InternalRel of R by XBOOLE_0:def_4; [z,x] in the InternalRel of R ~ by A3, XBOOLE_0:def_4; then A5: [x,z] in the InternalRel of R by RELAT_1:def_7; z in dom the InternalRel of R by A4, XTUPLE_0:def_12; then z = x by A2, A4, A5, RELAT_2:def_4; hence z in {x} by TARSKI:def_1; ::_thesis: verum end; assume z in {x} ; ::_thesis: z in Class ((EqRel R),x) then z = x by TARSKI:def_1; hence z in Class ((EqRel R),x) by EQREL_1:20; ::_thesis: verum end; hence Class ((EqRel R),x) = {x} by TARSKI:1; ::_thesis: verum end; theorem :: DICKSON:12 for R being Relation holds ( R = R \~ iff R is asymmetric ) proof let R be Relation; ::_thesis: ( R = R \~ iff R is asymmetric ) thus ( R = R \~ implies R is asymmetric ) ; ::_thesis: ( R is asymmetric implies R = R \~ ) assume R is asymmetric ; ::_thesis: R = R \~ then A1: R is_asymmetric_in field R by RELAT_2:def_13; now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_R_implies_[a,b]_in_R_\~_)_&_(_[a,b]_in_R_\~_implies_[a,b]_in_R_)_) let a, b be set ; ::_thesis: ( ( [a,b] in R implies [a,b] in R \~ ) & ( [a,b] in R \~ implies [a,b] in R ) ) hereby ::_thesis: ( [a,b] in R \~ implies [a,b] in R ) assume A2: [a,b] in R ; ::_thesis: [a,b] in R \~ then A3: a in field R by RELAT_1:15; b in field R by A2, RELAT_1:15; then not [b,a] in R by A1, A2, A3, RELAT_2:def_5; then not [a,b] in R ~ by RELAT_1:def_7; hence [a,b] in R \~ by A2, XBOOLE_0:def_5; ::_thesis: verum end; assume [a,b] in R \~ ; ::_thesis: [a,b] in R hence [a,b] in R ; ::_thesis: verum end; hence R = R \~ by RELAT_1:def_2; ::_thesis: verum end; theorem :: DICKSON:13 for R being Relation st R is transitive holds R \~ is transitive proof let R be Relation; ::_thesis: ( R is transitive implies R \~ is transitive ) assume R is transitive ; ::_thesis: R \~ is transitive then A1: R is_transitive_in field R by RELAT_2:def_16; now__::_thesis:_for_x,_y,_z_being_set_st_x_in_field_(R_\~)_&_y_in_field_(R_\~)_&_z_in_field_(R_\~)_&_[x,y]_in_R_\~_&_[y,z]_in_R_\~_holds_ [x,z]_in_R_\~ let x, y, z be set ; ::_thesis: ( x in field (R \~) & y in field (R \~) & z in field (R \~) & [x,y] in R \~ & [y,z] in R \~ implies [x,z] in R \~ ) assume that x in field (R \~) and y in field (R \~) and z in field (R \~) and A2: [x,y] in R \~ and A3: [y,z] in R \~ ; ::_thesis: [x,z] in R \~ A4: x in field R by A2, RELAT_1:15; A5: y in field R by A2, RELAT_1:15; A6: z in field R by A3, RELAT_1:15; then A7: [x,z] in R by A1, A2, A3, A4, A5, RELAT_2:def_8; not [x,y] in R ~ by A2, XBOOLE_0:def_5; then not [y,x] in R by RELAT_1:def_7; then not [z,x] in R by A1, A3, A4, A5, A6, RELAT_2:def_8; then not [x,z] in R ~ by RELAT_1:def_7; hence [x,z] in R \~ by A7, XBOOLE_0:def_5; ::_thesis: verum end; then R \~ is_transitive_in field (R \~) by RELAT_2:def_8; hence R \~ is transitive by RELAT_2:def_16; ::_thesis: verum end; theorem :: DICKSON:14 for R being Relation for a, b being set st R is antisymmetric holds ( [a,b] in R \~ iff ( [a,b] in R & a <> b ) ) proof let R be Relation; ::_thesis: for a, b being set st R is antisymmetric holds ( [a,b] in R \~ iff ( [a,b] in R & a <> b ) ) let a, b be set ; ::_thesis: ( R is antisymmetric implies ( [a,b] in R \~ iff ( [a,b] in R & a <> b ) ) ) assume R is antisymmetric ; ::_thesis: ( [a,b] in R \~ iff ( [a,b] in R & a <> b ) ) then A1: R is_antisymmetric_in field R by RELAT_2:def_12; A2: R \~ is_asymmetric_in field (R \~) by RELAT_2:def_13; hereby ::_thesis: ( [a,b] in R & a <> b implies [a,b] in R \~ ) assume A3: [a,b] in R \~ ; ::_thesis: ( [a,b] in R & a <> b ) hence [a,b] in R ; ::_thesis: a <> b now__::_thesis:_not_a_=_b assume A4: a = b ; ::_thesis: contradiction a in field (R \~) by A3, RELAT_1:15; hence contradiction by A2, A3, A4, RELAT_2:def_5; ::_thesis: verum end; hence a <> b ; ::_thesis: verum end; assume that A5: [a,b] in R and A6: a <> b ; ::_thesis: [a,b] in R \~ A7: a in field R by A5, RELAT_1:15; b in field R by A5, RELAT_1:15; then not [b,a] in R by A1, A5, A6, A7, RELAT_2:def_4; then not [a,b] in R ~ by RELAT_1:def_7; hence [a,b] in R \~ by A5, XBOOLE_0:def_5; ::_thesis: verum end; theorem Th15: :: DICKSON:15 for R being RelStr st R is well_founded holds R \~ is well_founded proof let R be RelStr ; ::_thesis: ( R is well_founded implies R \~ is well_founded ) assume A1: R is well_founded ; ::_thesis: R \~ is well_founded set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); set CR9 = the carrier of (R \~); A2: the InternalRel of R is_well_founded_in the carrier of R by A1, WELLFND1:def_2; now__::_thesis:_for_Y_being_set_st_Y_c=_the_carrier_of_(R_\~)_&_Y_<>_{}_holds_ ex_a_being_set_st_ (_a_in_Y_&_the_InternalRel_of_(R_\~)_-Seg_a_misses_Y_) let Y be set ; ::_thesis: ( Y c= the carrier of (R \~) & Y <> {} implies ex a being set st ( a in Y & the InternalRel of (R \~) -Seg a misses Y ) ) assume that A3: Y c= the carrier of (R \~) and A4: Y <> {} ; ::_thesis: ex a being set st ( a in Y & the InternalRel of (R \~) -Seg a misses Y ) consider a being set such that A5: a in Y and A6: the InternalRel of R -Seg a misses Y by A2, A3, A4, WELLORD1:def_3; A7: ( the InternalRel of R -Seg a) /\ Y = {} by A6, XBOOLE_0:def_7; take a = a; ::_thesis: ( a in Y & the InternalRel of (R \~) -Seg a misses Y ) thus a in Y by A5; ::_thesis: the InternalRel of (R \~) -Seg a misses Y now__::_thesis:_for_z_being_set_holds_not_z_in_(_the_InternalRel_of_(R_\~)_-Seg_a)_/\_Y assume ex z being set st z in ( the InternalRel of (R \~) -Seg a) /\ Y ; ::_thesis: contradiction then consider z being set such that A8: z in ( the InternalRel of (R \~) -Seg a) /\ Y ; A9: z in the InternalRel of (R \~) -Seg a by A8, XBOOLE_0:def_4; A10: z in Y by A8, XBOOLE_0:def_4; A11: z <> a by A9, WELLORD1:1; [z,a] in the InternalRel of (R \~) by A9, WELLORD1:1; then z in the InternalRel of R -Seg a by A11, WELLORD1:1; hence contradiction by A7, A10, XBOOLE_0:def_4; ::_thesis: verum end; then ( the InternalRel of (R \~) -Seg a) /\ Y = {} by XBOOLE_0:def_1; hence the InternalRel of (R \~) -Seg a misses Y by XBOOLE_0:def_7; ::_thesis: verum end; then the InternalRel of (R \~) is_well_founded_in the carrier of (R \~) by WELLORD1:def_3; hence R \~ is well_founded by WELLFND1:def_2; ::_thesis: verum end; theorem Th16: :: DICKSON:16 for R being RelStr st R \~ is well_founded & R is antisymmetric holds R is well_founded proof let R be RelStr ; ::_thesis: ( R \~ is well_founded & R is antisymmetric implies R is well_founded ) assume that A1: R \~ is well_founded and A2: R is antisymmetric ; ::_thesis: R is well_founded set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); A3: the InternalRel of R is_antisymmetric_in the carrier of R by A2, ORDERS_2:def_4; A4: the InternalRel of (R \~) is_well_founded_in the carrier of R by A1, WELLFND1:def_2; now__::_thesis:_for_Y_being_set_st_Y_c=_the_carrier_of_R_&_Y_<>_{}_holds_ ex_a_being_set_st_ (_a_in_Y_&_the_InternalRel_of_R_-Seg_a_misses_Y_) let Y be set ; ::_thesis: ( Y c= the carrier of R & Y <> {} implies ex a being set st ( a in Y & the InternalRel of R -Seg a misses Y ) ) assume that A5: Y c= the carrier of R and A6: Y <> {} ; ::_thesis: ex a being set st ( a in Y & the InternalRel of R -Seg a misses Y ) consider a being set such that A7: a in Y and A8: the InternalRel of (R \~) -Seg a misses Y by A4, A5, A6, WELLORD1:def_3; A9: ( the InternalRel of (R \~) -Seg a) /\ Y = {} by A8, XBOOLE_0:def_7; take a = a; ::_thesis: ( a in Y & the InternalRel of R -Seg a misses Y ) thus a in Y by A7; ::_thesis: the InternalRel of R -Seg a misses Y now__::_thesis:_not_(_the_InternalRel_of_R_-Seg_a)_/\_Y_<>_{} assume ( the InternalRel of R -Seg a) /\ Y <> {} ; ::_thesis: contradiction then consider z being set such that A10: z in ( the InternalRel of R -Seg a) /\ Y by XBOOLE_0:def_1; A11: z in the InternalRel of R -Seg a by A10, XBOOLE_0:def_4; A12: z in Y by A10, XBOOLE_0:def_4; A13: z <> a by A11, WELLORD1:1; A14: [z,a] in the InternalRel of R by A11, WELLORD1:1; then not [a,z] in the InternalRel of R by A3, A5, A7, A12, A13, RELAT_2:def_4; then not [z,a] in the InternalRel of R ~ by RELAT_1:def_7; then [z,a] in the InternalRel of R \ ( the InternalRel of R ~) by A14, XBOOLE_0:def_5; then z in the InternalRel of (R \~) -Seg a by A13, WELLORD1:1; hence contradiction by A9, A12, XBOOLE_0:def_4; ::_thesis: verum end; hence the InternalRel of R -Seg a misses Y by XBOOLE_0:def_7; ::_thesis: verum end; then the InternalRel of R is_well_founded_in the carrier of R by WELLORD1:def_3; hence R is well_founded by WELLFND1:def_2; ::_thesis: verum end; begin theorem Th17: :: DICKSON:17 for L being RelStr for N being set for x being Element of (L \~) holds ( x is_minimal_wrt N, the InternalRel of (L \~) iff ( x in N & ( for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ) ) ) proof let L be RelStr ; ::_thesis: for N being set for x being Element of (L \~) holds ( x is_minimal_wrt N, the InternalRel of (L \~) iff ( x in N & ( for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ) ) ) let N be set ; ::_thesis: for x being Element of (L \~) holds ( x is_minimal_wrt N, the InternalRel of (L \~) iff ( x in N & ( for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ) ) ) let x be Element of (L \~); ::_thesis: ( x is_minimal_wrt N, the InternalRel of (L \~) iff ( x in N & ( for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ) ) ) set IR = the InternalRel of L; set IR9 = the InternalRel of (L \~); hereby ::_thesis: ( x in N & ( for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ) implies x is_minimal_wrt N, the InternalRel of (L \~) ) assume A1: x is_minimal_wrt N, the InternalRel of (L \~) ; ::_thesis: ( x in N & ( for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ) ) hence x in N by WAYBEL_4:def_25; ::_thesis: for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L let y be Element of L; ::_thesis: ( y in N & [y,x] in the InternalRel of L implies [x,y] in the InternalRel of L ) assume A2: y in N ; ::_thesis: ( [y,x] in the InternalRel of L implies [x,y] in the InternalRel of L ) assume A3: [y,x] in the InternalRel of L ; ::_thesis: [x,y] in the InternalRel of L now__::_thesis:_[x,y]_in_the_InternalRel_of_L percases ( x = y or x <> y ) ; suppose x = y ; ::_thesis: [x,y] in the InternalRel of L hence [x,y] in the InternalRel of L by A3; ::_thesis: verum end; suppose x <> y ; ::_thesis: [x,y] in the InternalRel of L then not [y,x] in the InternalRel of (L \~) by A1, A2, WAYBEL_4:def_25; then [y,x] in the InternalRel of L ~ by A3, XBOOLE_0:def_5; hence [x,y] in the InternalRel of L by RELAT_1:def_7; ::_thesis: verum end; end; end; hence [x,y] in the InternalRel of L ; ::_thesis: verum end; assume that A4: x in N and A5: for y being Element of L st y in N & [y,x] in the InternalRel of L holds [x,y] in the InternalRel of L ; ::_thesis: x is_minimal_wrt N, the InternalRel of (L \~) now__::_thesis:_for_y_being_set_holds_ (_not_y_in_N_or_not_y_<>_x_or_not_[y,x]_in_the_InternalRel_of_(L_\~)_) assume ex y being set st ( y in N & y <> x & [y,x] in the InternalRel of (L \~) ) ; ::_thesis: contradiction then consider y being set such that A6: y in N and y <> x and A7: [y,x] in the InternalRel of (L \~) ; reconsider y9 = y as Element of L by A7, ZFMISC_1:87; A8: not [y,x] in the InternalRel of L ~ by A7, XBOOLE_0:def_5; ( [y9,x] in the InternalRel of L implies [x,y9] in the InternalRel of L ) by A5, A6; hence contradiction by A7, A8, RELAT_1:def_7; ::_thesis: verum end; hence x is_minimal_wrt N, the InternalRel of (L \~) by A4, WAYBEL_4:def_25; ::_thesis: verum end; theorem :: DICKSON:18 for R, S being non empty RelStr for m being Function of R,S st R is quasi_ordered & S is antisymmetric & S \~ is well_founded & ( for a, b being Element of R holds ( ( a <= b implies m . a <= m . b ) & ( m . a = m . b implies [a,b] in EqRel R ) ) ) holds R \~ is well_founded proof let R, S be non empty RelStr ; ::_thesis: for m being Function of R,S st R is quasi_ordered & S is antisymmetric & S \~ is well_founded & ( for a, b being Element of R holds ( ( a <= b implies m . a <= m . b ) & ( m . a = m . b implies [a,b] in EqRel R ) ) ) holds R \~ is well_founded let m be Function of R,S; ::_thesis: ( R is quasi_ordered & S is antisymmetric & S \~ is well_founded & ( for a, b being Element of R holds ( ( a <= b implies m . a <= m . b ) & ( m . a = m . b implies [a,b] in EqRel R ) ) ) implies R \~ is well_founded ) assume that A1: R is quasi_ordered and A2: S is antisymmetric and A3: S \~ is well_founded and A4: for a, b being Element of R holds ( ( a <= b implies m . a <= m . b ) & ( m . a = m . b implies [a,b] in EqRel R ) ) ; ::_thesis: R \~ is well_founded set IR = the InternalRel of R; set IS = the InternalRel of S; A5: the InternalRel of S is_antisymmetric_in the carrier of S by A2, ORDERS_2:def_4; now__::_thesis:_for_f_being_sequence_of_(R_\~)_holds_not_f_is_descending assume ex f being sequence of (R \~) st f is descending ; ::_thesis: contradiction then consider f being sequence of (R \~) such that A6: f is descending ; reconsider f9 = f as sequence of R ; deffunc H1( Element of NAT ) -> Element of the carrier of S = m . (f9 . $1); consider g9 being Function of NAT, the carrier of S such that A7: for k being Element of NAT holds g9 . k = H1(k) from FUNCT_2:sch_4(); reconsider g = g9 as sequence of (S \~) ; now__::_thesis:_for_n_being_Nat_holds_ (_g_._(n_+_1)_<>_g_._n_&_[(g_._(n_+_1)),(g_._n)]_in_the_InternalRel_of_(S_\~)_) let n be Nat; ::_thesis: ( g . (n + 1) <> g . n & [(g . (n + 1)),(g . n)] in the InternalRel of (S \~) ) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; A8: [(f . (n + 1)),(f . n)] in the InternalRel of (R \~) by A6, WELLFND1:def_6; A9: [(f . (n + 1)),(f . n)] in the InternalRel of R \ ( the InternalRel of R ~) by A6, WELLFND1:def_6; A10: not [(f . (n + 1)),(f . n)] in the InternalRel of R ~ by A8, XBOOLE_0:def_5; A11: g . n1 = m . (f . n1) by A7; A12: now__::_thesis:_not_g_._(n_+_1)_=_g_._n assume g . (n + 1) = g . n ; ::_thesis: contradiction then m . (f . (n + 1)) = m . (f . n) by A7, A11; then [(f9 . (n + 1)),(f9 . n)] in EqRel R by A4; then [(f . (n + 1)),(f . n)] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; hence contradiction by A10, XBOOLE_0:def_4; ::_thesis: verum end; hence g . (n + 1) <> g . n ; ::_thesis: [(g . (n + 1)),(g . n)] in the InternalRel of (S \~) reconsider fn1 = f . (n + 1) as Element of R ; reconsider fn = f . n as Element of R ; A13: fn1 <= fn by A9, ORDERS_2:def_5; A14: g9 . (n + 1) = m . fn1 by A7; g9 . n1 = m . fn by A7; then g9 . (n + 1) <= g9 . n by A4, A13, A14; then A15: [(g . (n + 1)),(g . n)] in the InternalRel of S by ORDERS_2:def_5; then not [(g . n),(g . (n + 1))] in the InternalRel of S by A5, A12, RELAT_2:def_4; then not [(g . (n + 1)),(g . n)] in the InternalRel of S ~ by RELAT_1:def_7; hence [(g . (n + 1)),(g . n)] in the InternalRel of (S \~) by A15, XBOOLE_0:def_5; ::_thesis: verum end; then g is descending by WELLFND1:def_6; hence contradiction by A3, WELLFND1:14; ::_thesis: verum end; hence R \~ is well_founded by WELLFND1:14; ::_thesis: verum end; definition let R be non empty RelStr ; let N be Subset of R; func min-classes N -> Subset-Family of R means :Def8: :: DICKSON:def 8 for x being set holds ( x in it iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ); existence ex b1 being Subset-Family of R st for x being set holds ( x in b1 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) proof set IR9 = the InternalRel of (R \~); set C = { (x /\ N) where x is Element of Class (EqRel R) : ex y being Element of (R \~) st ( x = Class ((EqRel R),y) & y is_minimal_wrt N, the InternalRel of (R \~) ) } ; now__::_thesis:_for_x_being_set_st_x_in__{__(x_/\_N)_where_x_is_Element_of_Class_(EqRel_R)_:_ex_y_being_Element_of_(R_\~)_st_ (_x_=_Class_((EqRel_R),y)_&_y_is_minimal_wrt_N,_the_InternalRel_of_(R_\~)_)__}__holds_ x_in_bool_the_carrier_of_R let x be set ; ::_thesis: ( x in { (x /\ N) where x is Element of Class (EqRel R) : ex y being Element of (R \~) st ( x = Class ((EqRel R),y) & y is_minimal_wrt N, the InternalRel of (R \~) ) } implies x in bool the carrier of R ) assume x in { (x /\ N) where x is Element of Class (EqRel R) : ex y being Element of (R \~) st ( x = Class ((EqRel R),y) & y is_minimal_wrt N, the InternalRel of (R \~) ) } ; ::_thesis: x in bool the carrier of R then ex xx being Element of Class (EqRel R) st ( x = xx /\ N & ex y being Element of (R \~) st ( xx = Class ((EqRel R),y) & y is_minimal_wrt N, the InternalRel of (R \~) ) ) ; hence x in bool the carrier of R ; ::_thesis: verum end; then reconsider C = { (x /\ N) where x is Element of Class (EqRel R) : ex y being Element of (R \~) st ( x = Class ((EqRel R),y) & y is_minimal_wrt N, the InternalRel of (R \~) ) } as Subset-Family of R by TARSKI:def_3; take C ; ::_thesis: for x being set holds ( x in C iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) let x be set ; ::_thesis: ( x in C iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) hereby ::_thesis: ( ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) implies x in C ) assume x in C ; ::_thesis: ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) then ex xx being Element of Class (EqRel R) st ( x = xx /\ N & ex y being Element of (R \~) st ( xx = Class ((EqRel R),y) & y is_minimal_wrt N, the InternalRel of (R \~) ) ) ; hence ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ; ::_thesis: verum end; given y being Element of (R \~) such that A1: y is_minimal_wrt N, the InternalRel of (R \~) and A2: x = (Class ((EqRel R),y)) /\ N ; ::_thesis: x in C reconsider y9 = y as Element of R ; EqClass ((EqRel R),y9) in Class (EqRel R) ; hence x in C by A1, A2; ::_thesis: verum end; uniqueness for b1, b2 being Subset-Family of R st ( for x being set holds ( x in b1 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) ) & ( for x being set holds ( x in b2 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) ) holds b1 = b2 proof set IR = the InternalRel of (R \~); let IT1, IT2 be Subset-Family of R; ::_thesis: ( ( for x being set holds ( x in IT1 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) ) & ( for x being set holds ( x in IT2 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) ) implies IT1 = IT2 ) assume that A3: for x being set holds ( x in IT1 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) and A4: for x being set holds ( x in IT2 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_x_being_set_holds_ (_(_x_in_IT1_implies_x_in_IT2_)_&_(_x_in_IT2_implies_x_in_IT1_)_) let x be set ; ::_thesis: ( ( x in IT1 implies x in IT2 ) & ( x in IT2 implies x in IT1 ) ) hereby ::_thesis: ( x in IT2 implies x in IT1 ) assume x in IT1 ; ::_thesis: x in IT2 then ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) by A3; hence x in IT2 by A4; ::_thesis: verum end; assume x in IT2 ; ::_thesis: x in IT1 then ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) by A4; hence x in IT1 by A3; ::_thesis: verum end; hence IT1 = IT2 by TARSKI:1; ::_thesis: verum end; end; :: deftheorem Def8 defines min-classes DICKSON:def_8_:_ for R being non empty RelStr for N being Subset of R for b3 being Subset-Family of R holds ( b3 = min-classes N iff for x being set holds ( x in b3 iff ex y being Element of (R \~) st ( y is_minimal_wrt N, the InternalRel of (R \~) & x = (Class ((EqRel R),y)) /\ N ) ) ); theorem Th19: :: DICKSON:19 for R being non empty RelStr for N being Subset of R for x being set st R is quasi_ordered & x in min-classes N holds for y being Element of (R \~) st y in x holds y is_minimal_wrt N, the InternalRel of (R \~) proof let R be non empty RelStr ; ::_thesis: for N being Subset of R for x being set st R is quasi_ordered & x in min-classes N holds for y being Element of (R \~) st y in x holds y is_minimal_wrt N, the InternalRel of (R \~) let N be Subset of R; ::_thesis: for x being set st R is quasi_ordered & x in min-classes N holds for y being Element of (R \~) st y in x holds y is_minimal_wrt N, the InternalRel of (R \~) let x be set ; ::_thesis: ( R is quasi_ordered & x in min-classes N implies for y being Element of (R \~) st y in x holds y is_minimal_wrt N, the InternalRel of (R \~) ) assume that A1: R is quasi_ordered and A2: x in min-classes N ; ::_thesis: for y being Element of (R \~) st y in x holds y is_minimal_wrt N, the InternalRel of (R \~) set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); let c be Element of (R \~); ::_thesis: ( c in x implies c is_minimal_wrt N, the InternalRel of (R \~) ) assume A3: c in x ; ::_thesis: c is_minimal_wrt N, the InternalRel of (R \~) consider b being Element of (R \~) such that A4: b is_minimal_wrt N, the InternalRel of (R \~) and A5: x = (Class ((EqRel R),b)) /\ N by A2, Def8; c in Class ((EqRel R),b) by A3, A5, XBOOLE_0:def_4; then [c,b] in EqRel R by EQREL_1:19; then [c,b] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; then A6: [c,b] in the InternalRel of R by XBOOLE_0:def_4; A7: now__::_thesis:_for_d_being_set_holds_ (_not_d_in_N_or_not_d_<>_c_or_not_[d,c]_in_the_InternalRel_of_(R_\~)_) assume ex d being set st ( d in N & d <> c & [d,c] in the InternalRel of (R \~) ) ; ::_thesis: contradiction then consider d being set such that A8: d in N and d <> c and A9: [d,c] in the InternalRel of (R \~) ; A10: not [d,c] in the InternalRel of R ~ by A9, XBOOLE_0:def_5; R is transitive by A1, Def3; then A11: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; then A12: [d,b] in the InternalRel of R by A6, A8, A9, RELAT_2:def_8; now__::_thesis:_not_[d,b]_in_the_InternalRel_of_R_~ assume [d,b] in the InternalRel of R ~ ; ::_thesis: contradiction then [b,d] in the InternalRel of R by RELAT_1:def_7; then [c,d] in the InternalRel of R by A6, A8, A11, RELAT_2:def_8; hence contradiction by A10, RELAT_1:def_7; ::_thesis: verum end; then A13: [d,b] in the InternalRel of (R \~) by A12, XBOOLE_0:def_5; b <> d by A6, A10, RELAT_1:def_7; hence contradiction by A4, A8, A13, WAYBEL_4:def_25; ::_thesis: verum end; c in N by A3, A5, XBOOLE_0:def_4; hence c is_minimal_wrt N, the InternalRel of (R \~) by A7, WAYBEL_4:def_25; ::_thesis: verum end; theorem Th20: :: DICKSON:20 for R being non empty RelStr holds ( R \~ is well_founded iff for N being Subset of R st N <> {} holds ex x being set st x in min-classes N ) proof let R be non empty RelStr ; ::_thesis: ( R \~ is well_founded iff for N being Subset of R st N <> {} holds ex x being set st x in min-classes N ) set CR = the carrier of R; set IR9 = the InternalRel of (R \~); set CR9 = the carrier of (R \~); hereby ::_thesis: ( ( for N being Subset of R st N <> {} holds ex x being set st x in min-classes N ) implies R \~ is well_founded ) assume R \~ is well_founded ; ::_thesis: for N being Subset of the carrier of R st N <> {} holds ex x being set st x in min-classes N then A1: the InternalRel of (R \~) is_well_founded_in the carrier of (R \~) by WELLFND1:def_2; let N be Subset of the carrier of R; ::_thesis: ( N <> {} implies ex x being set st x in min-classes N ) assume A2: N <> {} ; ::_thesis: ex x being set st x in min-classes N reconsider N9 = N as Subset of the carrier of (R \~) ; consider y being set such that A3: y in N9 and A4: the InternalRel of (R \~) -Seg y misses N9 by A1, A2, WELLORD1:def_3; A5: ( the InternalRel of (R \~) -Seg y) /\ N9 = {} by A4, XBOOLE_0:def_7; reconsider y = y as Element of (R \~) by A3; set x = (Class ((EqRel R),y)) /\ N; now__::_thesis:_for_z_being_set_holds_ (_not_z_in_N_or_not_z_<>_y_or_not_[z,y]_in_the_InternalRel_of_(R_\~)_) assume ex z being set st ( z in N & z <> y & [z,y] in the InternalRel of (R \~) ) ; ::_thesis: contradiction then consider z being set such that A6: z in N and A7: z <> y and A8: [z,y] in the InternalRel of (R \~) ; z in the InternalRel of (R \~) -Seg y by A7, A8, WELLORD1:1; hence contradiction by A5, A6, XBOOLE_0:def_4; ::_thesis: verum end; then y is_minimal_wrt N, the InternalRel of (R \~) by A3, WAYBEL_4:def_25; then (Class ((EqRel R),y)) /\ N in min-classes N by Def8; hence ex x being set st x in min-classes N ; ::_thesis: verum end; assume A9: for N being Subset of R st N <> {} holds ex x being set st x in min-classes N ; ::_thesis: R \~ is well_founded now__::_thesis:_for_N_being_set_st_N_c=_the_carrier_of_(R_\~)_&_N_<>_{}_holds_ ex_a9_being_set_st_ (_a9_in_N_&_the_InternalRel_of_(R_\~)_-Seg_a9_misses_N_) let N be set ; ::_thesis: ( N c= the carrier of (R \~) & N <> {} implies ex a9 being set st ( a9 in N & the InternalRel of (R \~) -Seg a9 misses N ) ) assume that A10: N c= the carrier of (R \~) and A11: N <> {} ; ::_thesis: ex a9 being set st ( a9 in N & the InternalRel of (R \~) -Seg a9 misses N ) reconsider N9 = N as Subset of R by A10; consider x being set such that A12: x in min-classes N9 by A9, A11; consider a being Element of (R \~) such that A13: a is_minimal_wrt N9, the InternalRel of (R \~) and x = (Class ((EqRel R),a)) /\ N9 by A12, Def8; reconsider a9 = a as set ; take a9 = a9; ::_thesis: ( a9 in N & the InternalRel of (R \~) -Seg a9 misses N ) thus a9 in N by A13, WAYBEL_4:def_25; ::_thesis: the InternalRel of (R \~) -Seg a9 misses N now__::_thesis:_not_(_the_InternalRel_of_(R_\~)_-Seg_a9)_/\_N_<>_{} assume ( the InternalRel of (R \~) -Seg a9) /\ N <> {} ; ::_thesis: contradiction then consider z being set such that A14: z in ( the InternalRel of (R \~) -Seg a9) /\ N by XBOOLE_0:def_1; A15: z in the InternalRel of (R \~) -Seg a9 by A14, XBOOLE_0:def_4; A16: z in N by A14, XBOOLE_0:def_4; then reconsider z = z as Element of (R \~) by A10; A17: z <> a9 by A15, WELLORD1:1; [z,a] in the InternalRel of (R \~) by A15, WELLORD1:1; hence contradiction by A13, A16, A17, WAYBEL_4:def_25; ::_thesis: verum end; hence the InternalRel of (R \~) -Seg a9 misses N by XBOOLE_0:def_7; ::_thesis: verum end; then the InternalRel of (R \~) is_well_founded_in the carrier of (R \~) by WELLORD1:def_3; hence R \~ is well_founded by WELLFND1:def_2; ::_thesis: verum end; theorem Th21: :: DICKSON:21 for R being non empty RelStr for N being Subset of R for y being Element of (R \~) st y is_minimal_wrt N, the InternalRel of (R \~) holds not min-classes N is empty proof let R be non empty RelStr ; ::_thesis: for N being Subset of R for y being Element of (R \~) st y is_minimal_wrt N, the InternalRel of (R \~) holds not min-classes N is empty let N be Subset of R; ::_thesis: for y being Element of (R \~) st y is_minimal_wrt N, the InternalRel of (R \~) holds not min-classes N is empty let y be Element of (R \~); ::_thesis: ( y is_minimal_wrt N, the InternalRel of (R \~) implies not min-classes N is empty ) assume A1: y is_minimal_wrt N, the InternalRel of (R \~) ; ::_thesis: not min-classes N is empty ex x being set st x = (Class ((EqRel R),y)) /\ N ; hence not min-classes N is empty by A1, Def8; ::_thesis: verum end; theorem Th22: :: DICKSON:22 for R being non empty RelStr for N being Subset of R for x being set st x in min-classes N holds not x is empty proof let R be non empty RelStr ; ::_thesis: for N being Subset of R for x being set st x in min-classes N holds not x is empty let N be Subset of R; ::_thesis: for x being set st x in min-classes N holds not x is empty let x be set ; ::_thesis: ( x in min-classes N implies not x is empty ) assume x in min-classes N ; ::_thesis: not x is empty then consider y being Element of (R \~) such that A1: y is_minimal_wrt N, the InternalRel of (R \~) and A2: x = (Class ((EqRel R),y)) /\ N by Def8; A3: y in N by A1, WAYBEL_4:def_25; y in Class ((EqRel R),y) by EQREL_1:20; hence not x is empty by A2, A3, XBOOLE_0:def_4; ::_thesis: verum end; theorem Th23: :: DICKSON:23 for R being non empty RelStr st R is quasi_ordered holds ( ( R is connected & R \~ is well_founded ) iff for N being non empty Subset of R holds card (min-classes N) = 1 ) proof let R be non empty RelStr ; ::_thesis: ( R is quasi_ordered implies ( ( R is connected & R \~ is well_founded ) iff for N being non empty Subset of R holds card (min-classes N) = 1 ) ) assume A1: R is quasi_ordered ; ::_thesis: ( ( R is connected & R \~ is well_founded ) iff for N being non empty Subset of R holds card (min-classes N) = 1 ) set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); hereby ::_thesis: ( ( for N being non empty Subset of R holds card (min-classes N) = 1 ) implies ( R is connected & R \~ is well_founded ) ) assume that A2: R is connected and A3: R \~ is well_founded ; ::_thesis: for N being non empty Subset of the carrier of R holds card (min-classes N) = 1 let N be non empty Subset of the carrier of R; ::_thesis: card (min-classes N) = 1 consider x being set such that A4: x in min-classes N by A3, Th20; consider a being Element of (R \~) such that A5: a is_minimal_wrt N, the InternalRel of (R \~) and A6: x = (Class ((EqRel R),a)) /\ N by A4, Def8; A7: a in N by A5, WAYBEL_4:def_25; now__::_thesis:_for_y_being_set_holds_ (_(_y_in_min-classes_N_implies_y_in_{x}_)_&_(_y_in_{x}_implies_y_in_min-classes_N_)_) let y be set ; ::_thesis: ( ( y in min-classes N implies y in {x} ) & ( y in {x} implies y in min-classes N ) ) hereby ::_thesis: ( y in {x} implies y in min-classes N ) assume y in min-classes N ; ::_thesis: y in {x} then consider b being Element of (R \~) such that A8: b is_minimal_wrt N, the InternalRel of (R \~) and A9: y = (Class ((EqRel R),b)) /\ N by Def8; A10: b in N by A8, WAYBEL_4:def_25; reconsider a9 = a as Element of R ; reconsider b9 = b as Element of R ; A11: now__::_thesis:_(_[a,b]_in_the_InternalRel_of_R_&_[b,a]_in_the_InternalRel_of_R_) percases ( a9 <= b9 or b9 <= a9 ) by A2, WAYBEL_0:def_29; supposeA12: a9 <= b9 ; ::_thesis: ( [a,b] in the InternalRel of R & [b,a] in the InternalRel of R ) then A13: [a,b] in the InternalRel of R by ORDERS_2:def_5; now__::_thesis:_[b,a]_in_the_InternalRel_of_R percases ( a = b or a <> b ) ; suppose a = b ; ::_thesis: [b,a] in the InternalRel of R hence [b,a] in the InternalRel of R by A12, ORDERS_2:def_5; ::_thesis: verum end; supposeA14: a <> b ; ::_thesis: [b,a] in the InternalRel of R now__::_thesis:_[b,a]_in_the_InternalRel_of_R assume not [b,a] in the InternalRel of R ; ::_thesis: contradiction then not [a,b] in the InternalRel of R ~ by RELAT_1:def_7; then [a,b] in the InternalRel of R \ ( the InternalRel of R ~) by A13, XBOOLE_0:def_5; hence contradiction by A7, A8, A14, WAYBEL_4:def_25; ::_thesis: verum end; hence [b,a] in the InternalRel of R ; ::_thesis: verum end; end; end; hence ( [a,b] in the InternalRel of R & [b,a] in the InternalRel of R ) by A12, ORDERS_2:def_5; ::_thesis: verum end; supposeA15: b9 <= a9 ; ::_thesis: ( [a,b] in the InternalRel of R & [b,a] in the InternalRel of R ) then A16: [b,a] in the InternalRel of R by ORDERS_2:def_5; now__::_thesis:_[a,b]_in_the_InternalRel_of_R percases ( a = b or a <> b ) ; suppose a = b ; ::_thesis: [a,b] in the InternalRel of R hence [a,b] in the InternalRel of R by A15, ORDERS_2:def_5; ::_thesis: verum end; supposeA17: a <> b ; ::_thesis: [a,b] in the InternalRel of R now__::_thesis:_[a,b]_in_the_InternalRel_of_R assume not [a,b] in the InternalRel of R ; ::_thesis: contradiction then not [b,a] in the InternalRel of R ~ by RELAT_1:def_7; then [b,a] in the InternalRel of R \ ( the InternalRel of R ~) by A16, XBOOLE_0:def_5; hence contradiction by A5, A10, A17, WAYBEL_4:def_25; ::_thesis: verum end; hence [a,b] in the InternalRel of R ; ::_thesis: verum end; end; end; hence ( [a,b] in the InternalRel of R & [b,a] in the InternalRel of R ) by A15, ORDERS_2:def_5; ::_thesis: verum end; end; end; then [b,a] in the InternalRel of R ~ by RELAT_1:def_7; then [b,a] in the InternalRel of R /\ ( the InternalRel of R ~) by A11, XBOOLE_0:def_4; then [b,a] in EqRel R by A1, Def4; then b in Class ((EqRel R),a) by EQREL_1:19; then Class ((EqRel R),b) = Class ((EqRel R),a) by EQREL_1:23; hence y in {x} by A6, A9, TARSKI:def_1; ::_thesis: verum end; assume y in {x} ; ::_thesis: y in min-classes N then y = (Class ((EqRel R),a)) /\ N by A6, TARSKI:def_1; hence y in min-classes N by A5, Def8; ::_thesis: verum end; then min-classes N = {x} by TARSKI:1; hence card (min-classes N) = 1 by CARD_1:30; ::_thesis: verum end; assume A18: for N being non empty Subset of R holds card (min-classes N) = 1 ; ::_thesis: ( R is connected & R \~ is well_founded ) now__::_thesis:_for_a,_b_being_Element_of_R_st_not_a_<=_b_holds_ a_>=_b let a, b be Element of R; ::_thesis: ( not a <= b implies a >= b ) assume that A19: not a <= b and A20: not a >= b ; ::_thesis: contradiction A21: not [a,b] in the InternalRel of R by A19, ORDERS_2:def_5; then A22: not [b,a] in the InternalRel of R ~ by RELAT_1:def_7; not [b,a] in the InternalRel of R by A20, ORDERS_2:def_5; then A23: not [a,b] in the InternalRel of R ~ by RELAT_1:def_7; set N9 = {a,b}; set MCN = {{a},{b}}; now__::_thesis:_for_x_being_set_holds_ (_(_x_in_min-classes_{a,b}_implies_x_in_{{a},{b}}_)_&_(_x_in_{{a},{b}}_implies_x_in_min-classes_{a,b}_)_) let x be set ; ::_thesis: ( ( x in min-classes {a,b} implies x in {{a},{b}} ) & ( x in {{a},{b}} implies b1 in min-classes {a,b} ) ) hereby ::_thesis: ( x in {{a},{b}} implies b1 in min-classes {a,b} ) assume x in min-classes {a,b} ; ::_thesis: x in {{a},{b}} then consider y being Element of (R \~) such that A24: y is_minimal_wrt {a,b}, the InternalRel of (R \~) and A25: x = (Class ((EqRel R),y)) /\ {a,b} by Def8; A26: y in {a,b} by A24, WAYBEL_4:def_25; percases ( y = a or y = b ) by A26, TARSKI:def_2; supposeA27: y = a ; ::_thesis: x in {{a},{b}} now__::_thesis:_for_z_being_set_holds_ (_(_z_in_x_implies_z_in_{a}_)_&_(_z_in_{a}_implies_z_in_x_)_) let z be set ; ::_thesis: ( ( z in x implies z in {a} ) & ( z in {a} implies z in x ) ) hereby ::_thesis: ( z in {a} implies z in x ) assume A28: z in x ; ::_thesis: z in {a} then A29: z in Class ((EqRel R),a) by A25, A27, XBOOLE_0:def_4; A30: z in {a,b} by A25, A28, XBOOLE_0:def_4; now__::_thesis:_z_in_{a} percases ( z = a or z = b ) by A30, TARSKI:def_2; suppose z = a ; ::_thesis: z in {a} hence z in {a} by TARSKI:def_1; ::_thesis: verum end; suppose z = b ; ::_thesis: z in {a} then [b,a] in EqRel R by A29, EQREL_1:19; then [b,a] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; hence z in {a} by A22, XBOOLE_0:def_4; ::_thesis: verum end; end; end; hence z in {a} ; ::_thesis: verum end; assume z in {a} ; ::_thesis: z in x then A31: z = a by TARSKI:def_1; then A32: z in {a,b} by TARSKI:def_2; z in Class ((EqRel R),a) by A31, EQREL_1:20; hence z in x by A25, A27, A32, XBOOLE_0:def_4; ::_thesis: verum end; then x = {a} by TARSKI:1; hence x in {{a},{b}} by TARSKI:def_2; ::_thesis: verum end; supposeA33: y = b ; ::_thesis: x in {{a},{b}} now__::_thesis:_for_z_being_set_holds_ (_(_z_in_x_implies_z_in_{b}_)_&_(_z_in_{b}_implies_z_in_x_)_) let z be set ; ::_thesis: ( ( z in x implies z in {b} ) & ( z in {b} implies z in x ) ) hereby ::_thesis: ( z in {b} implies z in x ) assume A34: z in x ; ::_thesis: z in {b} then A35: z in Class ((EqRel R),b) by A25, A33, XBOOLE_0:def_4; A36: z in {a,b} by A25, A34, XBOOLE_0:def_4; now__::_thesis:_z_in_{b} percases ( z = b or z = a ) by A36, TARSKI:def_2; suppose z = b ; ::_thesis: z in {b} hence z in {b} by TARSKI:def_1; ::_thesis: verum end; suppose z = a ; ::_thesis: z in {b} then [a,b] in EqRel R by A35, EQREL_1:19; then [a,b] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; hence z in {b} by A21, XBOOLE_0:def_4; ::_thesis: verum end; end; end; hence z in {b} ; ::_thesis: verum end; assume z in {b} ; ::_thesis: z in x then A37: z = b by TARSKI:def_1; then A38: z in {a,b} by TARSKI:def_2; z in Class ((EqRel R),b) by A37, EQREL_1:20; hence z in x by A25, A33, A38, XBOOLE_0:def_4; ::_thesis: verum end; then x = {b} by TARSKI:1; hence x in {{a},{b}} by TARSKI:def_2; ::_thesis: verum end; end; end; assume A39: x in {{a},{b}} ; ::_thesis: b1 in min-classes {a,b} percases ( x = {a} or x = {b} ) by A39, TARSKI:def_2; supposeA40: x = {a} ; ::_thesis: b1 in min-classes {a,b} now__::_thesis:_ex_a9_being_Element_of_(R_\~)_st_ (_a9_is_minimal_wrt_{a,b},_the_InternalRel_of_(R_\~)_&_x_=_(Class_((EqRel_R),a9))_/\_{a,b}_) reconsider a9 = a as Element of (R \~) ; take a9 = a9; ::_thesis: ( a9 is_minimal_wrt {a,b}, the InternalRel of (R \~) & x = (Class ((EqRel R),a9)) /\ {a,b} ) A41: a9 in {a,b} by TARSKI:def_2; now__::_thesis:_for_y_being_set_holds_ (_not_y_in_{a,b}_or_not_y_<>_a9_or_not_[y,a9]_in_the_InternalRel_of_(R_\~)_) assume ex y being set st ( y in {a,b} & y <> a9 & [y,a9] in the InternalRel of (R \~) ) ; ::_thesis: contradiction then consider y being set such that A42: y in {a,b} and A43: y <> a9 and A44: [y,a9] in the InternalRel of (R \~) ; y = b by A42, A43, TARSKI:def_2; hence contradiction by A20, A44, ORDERS_2:def_5; ::_thesis: verum end; hence a9 is_minimal_wrt {a,b}, the InternalRel of (R \~) by A41, WAYBEL_4:def_25; ::_thesis: x = (Class ((EqRel R),a9)) /\ {a,b} now__::_thesis:_for_z_being_set_holds_ (_(_z_in_x_implies_z_in_(Class_((EqRel_R),a))_/\_{a,b}_)_&_(_z_in_(Class_((EqRel_R),a))_/\_{a,b}_implies_z_in_x_)_) let z be set ; ::_thesis: ( ( z in x implies z in (Class ((EqRel R),a)) /\ {a,b} ) & ( z in (Class ((EqRel R),a)) /\ {a,b} implies b1 in x ) ) hereby ::_thesis: ( z in (Class ((EqRel R),a)) /\ {a,b} implies b1 in x ) assume z in x ; ::_thesis: z in (Class ((EqRel R),a)) /\ {a,b} then A45: z = a by A40, TARSKI:def_1; then z in Class ((EqRel R),a) by EQREL_1:20; hence z in (Class ((EqRel R),a)) /\ {a,b} by A41, A45, XBOOLE_0:def_4; ::_thesis: verum end; assume A46: z in (Class ((EqRel R),a)) /\ {a,b} ; ::_thesis: b1 in x then A47: z in Class ((EqRel R),a) by XBOOLE_0:def_4; A48: z in {a,b} by A46, XBOOLE_0:def_4; percases ( z = a or z = b ) by A48, TARSKI:def_2; suppose z = a ; ::_thesis: b1 in x hence z in x by A40, TARSKI:def_1; ::_thesis: verum end; suppose z = b ; ::_thesis: b1 in x then [b,a] in EqRel R by A47, EQREL_1:19; then [b,a] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; hence z in x by A22, XBOOLE_0:def_4; ::_thesis: verum end; end; end; hence x = (Class ((EqRel R),a9)) /\ {a,b} by TARSKI:1; ::_thesis: verum end; hence x in min-classes {a,b} by Def8; ::_thesis: verum end; supposeA49: x = {b} ; ::_thesis: b1 in min-classes {a,b} now__::_thesis:_ex_b9_being_Element_of_(R_\~)_st_ (_b9_is_minimal_wrt_{a,b},_the_InternalRel_of_(R_\~)_&_x_=_(Class_((EqRel_R),b9))_/\_{a,b}_) reconsider b9 = b as Element of (R \~) ; take b9 = b9; ::_thesis: ( b9 is_minimal_wrt {a,b}, the InternalRel of (R \~) & x = (Class ((EqRel R),b9)) /\ {a,b} ) A50: b9 in {a,b} by TARSKI:def_2; now__::_thesis:_for_y_being_set_holds_ (_not_y_in_{a,b}_or_not_y_<>_b9_or_not_[y,b9]_in_the_InternalRel_of_(R_\~)_) assume ex y being set st ( y in {a,b} & y <> b9 & [y,b9] in the InternalRel of (R \~) ) ; ::_thesis: contradiction then consider y being set such that A51: y in {a,b} and A52: y <> b9 and A53: [y,b9] in the InternalRel of (R \~) ; y = a by A51, A52, TARSKI:def_2; hence contradiction by A19, A53, ORDERS_2:def_5; ::_thesis: verum end; hence b9 is_minimal_wrt {a,b}, the InternalRel of (R \~) by A50, WAYBEL_4:def_25; ::_thesis: x = (Class ((EqRel R),b9)) /\ {a,b} now__::_thesis:_for_z_being_set_holds_ (_(_z_in_x_implies_z_in_(Class_((EqRel_R),b))_/\_{a,b}_)_&_(_z_in_(Class_((EqRel_R),b))_/\_{a,b}_implies_z_in_x_)_) let z be set ; ::_thesis: ( ( z in x implies z in (Class ((EqRel R),b)) /\ {a,b} ) & ( z in (Class ((EqRel R),b)) /\ {a,b} implies b1 in x ) ) hereby ::_thesis: ( z in (Class ((EqRel R),b)) /\ {a,b} implies b1 in x ) assume z in x ; ::_thesis: z in (Class ((EqRel R),b)) /\ {a,b} then A54: z = b by A49, TARSKI:def_1; then z in Class ((EqRel R),b) by EQREL_1:20; hence z in (Class ((EqRel R),b)) /\ {a,b} by A50, A54, XBOOLE_0:def_4; ::_thesis: verum end; assume A55: z in (Class ((EqRel R),b)) /\ {a,b} ; ::_thesis: b1 in x then A56: z in Class ((EqRel R),b) by XBOOLE_0:def_4; A57: z in {a,b} by A55, XBOOLE_0:def_4; percases ( z = b or z = a ) by A57, TARSKI:def_2; suppose z = b ; ::_thesis: b1 in x hence z in x by A49, TARSKI:def_1; ::_thesis: verum end; suppose z = a ; ::_thesis: b1 in x then [a,b] in EqRel R by A56, EQREL_1:19; then [a,b] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; hence z in x by A23, XBOOLE_0:def_4; ::_thesis: verum end; end; end; hence x = (Class ((EqRel R),b9)) /\ {a,b} by TARSKI:1; ::_thesis: verum end; hence x in min-classes {a,b} by Def8; ::_thesis: verum end; end; end; then min-classes {a,b} = {{a},{b}} by TARSKI:1; then A58: card {{a},{b}} = 1 by A18; now__::_thesis:_not_{a}_=_{b} assume {a} = {b} ; ::_thesis: contradiction then A59: a = b by ZFMISC_1:3; R is reflexive by A1, Def3; then the InternalRel of R is_reflexive_in the carrier of R by ORDERS_2:def_2; hence contradiction by A21, A59, RELAT_2:def_1; ::_thesis: verum end; hence contradiction by A58, CARD_2:57; ::_thesis: verum end; hence R is connected by WAYBEL_0:def_29; ::_thesis: R \~ is well_founded now__::_thesis:_for_N_being_Subset_of_R_st_N_<>_{}_holds_ ex_x_being_set_st_x_in_min-classes_N let N be Subset of R; ::_thesis: ( N <> {} implies ex x being set st x in min-classes N ) assume N <> {} ; ::_thesis: ex x being set st x in min-classes N then card (min-classes N) = 1 by A18; then min-classes N <> {} ; hence ex x being set st x in min-classes N by XBOOLE_0:def_1; ::_thesis: verum end; hence R \~ is well_founded by Th20; ::_thesis: verum end; theorem :: DICKSON:24 for R being non empty Poset holds ( the InternalRel of R well_orders the carrier of R iff for N being non empty Subset of R holds card (min-classes N) = 1 ) proof let R be non empty Poset; ::_thesis: ( the InternalRel of R well_orders the carrier of R iff for N being non empty Subset of R holds card (min-classes N) = 1 ) set IR = the InternalRel of R; set CR = the carrier of R; A1: R is quasi_ordered by Def3; hereby ::_thesis: ( ( for N being non empty Subset of R holds card (min-classes N) = 1 ) implies the InternalRel of R well_orders the carrier of R ) assume A2: the InternalRel of R well_orders the carrier of R ; ::_thesis: for N being non empty Subset of R holds card (min-classes N) = 1 then A3: the InternalRel of R is_reflexive_in the carrier of R by WELLORD1:def_5; A4: the InternalRel of R is_connected_in the carrier of R by A2, WELLORD1:def_5; A5: the InternalRel of R is_well_founded_in the carrier of R by A2, WELLORD1:def_5; the InternalRel of R is_strongly_connected_in the carrier of R by A3, A4, ORDERS_1:7; then A6: R is connected by Th5; R is well_founded by A5, WELLFND1:def_2; then R \~ is well_founded by Th15; hence for N being non empty Subset of R holds card (min-classes N) = 1 by A1, A6, Th23; ::_thesis: verum end; assume A7: for N being non empty Subset of R holds card (min-classes N) = 1 ; ::_thesis: the InternalRel of R well_orders the carrier of R then A8: R is connected by A1, Th23; A9: R \~ is well_founded by A1, A7, Th23; A10: the InternalRel of R is_strongly_connected_in the carrier of R by A8, Th5; A11: R is well_founded by A9, Th16; A12: the InternalRel of R is_reflexive_in the carrier of R by ORDERS_2:def_2; A13: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; A14: the InternalRel of R is_antisymmetric_in the carrier of R by ORDERS_2:def_4; A15: the InternalRel of R is_connected_in the carrier of R by A10, ORDERS_1:7; the InternalRel of R is_well_founded_in the carrier of R by A11, WELLFND1:def_2; hence the InternalRel of R well_orders the carrier of R by A12, A13, A14, A15, WELLORD1:def_5; ::_thesis: verum end; definition let R be RelStr ; let N be Subset of R; let B be set ; predB is_Dickson-basis_of N,R means :Def9: :: DICKSON:def 9 ( B c= N & ( for a being Element of R st a in N holds ex b being Element of R st ( b in B & b <= a ) ) ); end; :: deftheorem Def9 defines is_Dickson-basis_of DICKSON:def_9_:_ for R being RelStr for N being Subset of R for B being set holds ( B is_Dickson-basis_of N,R iff ( B c= N & ( for a being Element of R st a in N holds ex b being Element of R st ( b in B & b <= a ) ) ) ); theorem Th25: :: DICKSON:25 for R being RelStr holds {} is_Dickson-basis_of {} the carrier of R,R proof let R be RelStr ; ::_thesis: {} is_Dickson-basis_of {} the carrier of R,R set N = {} the carrier of R; thus {} c= {} the carrier of R ; :: according to DICKSON:def_9 ::_thesis: for a being Element of R st a in {} the carrier of R holds ex b being Element of R st ( b in {} & b <= a ) thus for a being Element of R st a in {} the carrier of R holds ex b being Element of R st ( b in {} & b <= a ) ; ::_thesis: verum end; theorem Th26: :: DICKSON:26 for R being non empty RelStr for N being non empty Subset of R for B being set st B is_Dickson-basis_of N,R holds not B is empty proof let R be non empty RelStr ; ::_thesis: for N being non empty Subset of R for B being set st B is_Dickson-basis_of N,R holds not B is empty let N be non empty Subset of R; ::_thesis: for B being set st B is_Dickson-basis_of N,R holds not B is empty let B be set ; ::_thesis: ( B is_Dickson-basis_of N,R implies not B is empty ) assume A1: B is_Dickson-basis_of N,R ; ::_thesis: not B is empty set a = the Element of N; ex b being Element of R st ( b in B & b <= the Element of N ) by A1, Def9; hence not B is empty ; ::_thesis: verum end; definition let R be RelStr ; attrR is Dickson means :Def10: :: DICKSON:def 10 for N being Subset of R ex B being set st ( B is_Dickson-basis_of N,R & B is finite ); end; :: deftheorem Def10 defines Dickson DICKSON:def_10_:_ for R being RelStr holds ( R is Dickson iff for N being Subset of R ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) ); theorem Th27: :: DICKSON:27 for R being non empty RelStr st R \~ is well_founded & R is connected holds R is Dickson proof let R be non empty RelStr ; ::_thesis: ( R \~ is well_founded & R is connected implies R is Dickson ) assume that A1: R \~ is well_founded and A2: R is connected ; ::_thesis: R is Dickson set IR9 = the InternalRel of (R \~); set CR9 = the carrier of (R \~); set IR = the InternalRel of R; set CR = the carrier of R; let N be Subset of the carrier of R; :: according to DICKSON:def_10 ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) percases ( N = {} the carrier of R or N <> {} the carrier of R ) ; supposeA3: N = {} the carrier of R ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) take B = {} ; ::_thesis: ( B is_Dickson-basis_of N,R & B is finite ) thus B is_Dickson-basis_of N,R by A3, Th25; ::_thesis: B is finite thus B is finite ; ::_thesis: verum end; supposeA4: N <> {} the carrier of R ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) the InternalRel of (R \~) is_well_founded_in the carrier of (R \~) by A1, WELLFND1:def_2; then consider b being set such that A5: b in N and A6: the InternalRel of (R \~) -Seg b misses N by A4, WELLORD1:def_3; A7: ( the InternalRel of (R \~) -Seg b) /\ N = {} by A6, XBOOLE_0:def_7; take B = {b}; ::_thesis: ( B is_Dickson-basis_of N,R & B is finite ) reconsider b = b as Element of N by A5; thus B is_Dickson-basis_of N,R ::_thesis: B is finite proof {b} is Subset of N by A4, SUBSET_1:33; hence B c= N ; :: according to DICKSON:def_9 ::_thesis: for a being Element of R st a in N holds ex b being Element of R st ( b in B & b <= a ) let a be Element of R; ::_thesis: ( a in N implies ex b being Element of R st ( b in B & b <= a ) ) assume A8: a in N ; ::_thesis: ex b being Element of R st ( b in B & b <= a ) reconsider b = b as Element of R by A5; take b ; ::_thesis: ( b in B & b <= a ) thus b in B by TARSKI:def_1; ::_thesis: b <= a percases ( b <= a or a <= b ) by A2, WAYBEL_0:def_29; suppose b <= a ; ::_thesis: b <= a hence b <= a ; ::_thesis: verum end; supposeA9: a <= b ; ::_thesis: b <= a then A10: [a,b] in the InternalRel of R by ORDERS_2:def_5; now__::_thesis:_b_<=_a percases ( a = b or not a = b ) ; suppose a = b ; ::_thesis: b <= a hence b <= a by A9; ::_thesis: verum end; supposeA11: not a = b ; ::_thesis: b <= a now__::_thesis:_b_<=_a percases ( [a,b] in the InternalRel of (R \~) or not [a,b] in the InternalRel of (R \~) ) ; suppose [a,b] in the InternalRel of (R \~) ; ::_thesis: b <= a then a in the InternalRel of (R \~) -Seg b by A11, WELLORD1:1; hence b <= a by A7, A8, XBOOLE_0:def_4; ::_thesis: verum end; suppose not [a,b] in the InternalRel of (R \~) ; ::_thesis: b <= a then [a,b] in the InternalRel of R ~ by A10, XBOOLE_0:def_5; then [b,a] in the InternalRel of R by RELAT_1:def_7; hence b <= a by ORDERS_2:def_5; ::_thesis: verum end; end; end; hence b <= a ; ::_thesis: verum end; end; end; hence b <= a ; ::_thesis: verum end; end; end; thus B is finite ; ::_thesis: verum end; end; end; theorem Th28: :: DICKSON:28 for R, S being RelStr st the InternalRel of R c= the InternalRel of S & R is Dickson & the carrier of R = the carrier of S holds S is Dickson proof let r, s be RelStr ; ::_thesis: ( the InternalRel of r c= the InternalRel of s & r is Dickson & the carrier of r = the carrier of s implies s is Dickson ) assume that A1: the InternalRel of r c= the InternalRel of s and A2: r is Dickson and A3: the carrier of r = the carrier of s ; ::_thesis: s is Dickson let N be Subset of s; :: according to DICKSON:def_10 ::_thesis: ex B being set st ( B is_Dickson-basis_of N,s & B is finite ) reconsider N9 = N as Subset of r by A3; consider B being set such that A4: B is_Dickson-basis_of N9,r and A5: B is finite by A2, Def10; take B ; ::_thesis: ( B is_Dickson-basis_of N,s & B is finite ) thus B c= N by A4, Def9; :: according to DICKSON:def_9 ::_thesis: ( ( for a being Element of s st a in N holds ex b being Element of s st ( b in B & b <= a ) ) & B is finite ) hereby ::_thesis: B is finite let a be Element of s; ::_thesis: ( a in N implies ex b9 being Element of s st ( b9 in B & b9 <= a ) ) assume A6: a in N ; ::_thesis: ex b9 being Element of s st ( b9 in B & b9 <= a ) reconsider a9 = a as Element of r by A3; consider b being Element of r such that A7: b in B and A8: b <= a9 by A4, A6, Def9; reconsider b9 = b as Element of s by A3; take b9 = b9; ::_thesis: ( b9 in B & b9 <= a ) [b,a9] in the InternalRel of r by A8, ORDERS_2:def_5; hence ( b9 in B & b9 <= a ) by A1, A7, ORDERS_2:def_5; ::_thesis: verum end; thus B is finite by A5; ::_thesis: verum end; definition let f be Function; let b be set ; assume that B1: dom f = NAT and B2: b in rng f ; funcf mindex b -> Element of NAT means :Def11: :: DICKSON:def 11 ( f . it = b & ( for i being Element of NAT st f . i = b holds it <= i ) ); existence ex b1 being Element of NAT st ( f . b1 = b & ( for i being Element of NAT st f . i = b holds b1 <= i ) ) proof set N = { i where i is Element of NAT : f . i = b } ; consider x being set such that A1: x in NAT and A2: f . x = b by B1, B2, FUNCT_1:def_3; reconsider x = x as Element of NAT by A1; A3: x in { i where i is Element of NAT : f . i = b } by A2; now__::_thesis:_for_x_being_set_st_x_in__{__i_where_i_is_Element_of_NAT_:_f_._i_=_b__}__holds_ x_in_NAT let x be set ; ::_thesis: ( x in { i where i is Element of NAT : f . i = b } implies x in NAT ) assume x in { i where i is Element of NAT : f . i = b } ; ::_thesis: x in NAT then ex i being Element of NAT st ( x = i & f . i = b ) ; hence x in NAT ; ::_thesis: verum end; then reconsider N = { i where i is Element of NAT : f . i = b } as non empty Subset of NAT by A3, TARSKI:def_3; take I = min N; ::_thesis: ( f . I = b & ( for i being Element of NAT st f . i = b holds I <= i ) ) I in N by XXREAL_2:def_7; then ex II being Element of NAT st ( II = I & f . II = b ) ; hence f . I = b ; ::_thesis: for i being Element of NAT st f . i = b holds I <= i let i be Element of NAT ; ::_thesis: ( f . i = b implies I <= i ) assume f . i = b ; ::_thesis: I <= i then i in N ; hence I <= i by XXREAL_2:def_7; ::_thesis: verum end; uniqueness for b1, b2 being Element of NAT st f . b1 = b & ( for i being Element of NAT st f . i = b holds b1 <= i ) & f . b2 = b & ( for i being Element of NAT st f . i = b holds b2 <= i ) holds b1 = b2 proof let IT1, IT2 be Element of NAT ; ::_thesis: ( f . IT1 = b & ( for i being Element of NAT st f . i = b holds IT1 <= i ) & f . IT2 = b & ( for i being Element of NAT st f . i = b holds IT2 <= i ) implies IT1 = IT2 ) assume that A4: f . IT1 = b and A5: for i being Element of NAT st f . i = b holds IT1 <= i and A6: f . IT2 = b and A7: for i being Element of NAT st f . i = b holds IT2 <= i ; ::_thesis: IT1 = IT2 assume A8: IT1 <> IT2 ; ::_thesis: contradiction percases ( IT1 < IT2 or IT1 > IT2 ) by A8, XXREAL_0:1; suppose IT1 < IT2 ; ::_thesis: contradiction hence contradiction by A4, A7; ::_thesis: verum end; suppose IT1 > IT2 ; ::_thesis: contradiction hence contradiction by A5, A6; ::_thesis: verum end; end; end; end; :: deftheorem Def11 defines mindex DICKSON:def_11_:_ for f being Function for b being set st dom f = NAT & b in rng f holds for b3 being Element of NAT holds ( b3 = f mindex b iff ( f . b3 = b & ( for i being Element of NAT st f . i = b holds b3 <= i ) ) ); definition let R be non empty 1-sorted ; let f be sequence of R; let b be set ; let m be Element of NAT ; assume A1: ex j being Element of NAT st ( m < j & f . j = b ) ; funcf mindex (b,m) -> Element of NAT means :Def12: :: DICKSON:def 12 ( f . it = b & m < it & ( for i being Element of NAT st m < i & f . i = b holds it <= i ) ); existence ex b1 being Element of NAT st ( f . b1 = b & m < b1 & ( for i being Element of NAT st m < i & f . i = b holds b1 <= i ) ) proof set N = { i where i is Element of NAT : ( m < i & f . i = b ) } ; consider j being Element of NAT such that A2: m < j and A3: f . j = b by A1; A4: j in { i where i is Element of NAT : ( m < i & f . i = b ) } by A2, A3; now__::_thesis:_for_x_being_set_st_x_in__{__i_where_i_is_Element_of_NAT_:_(_m_<_i_&_f_._i_=_b_)__}__holds_ x_in_NAT let x be set ; ::_thesis: ( x in { i where i is Element of NAT : ( m < i & f . i = b ) } implies x in NAT ) assume x in { i where i is Element of NAT : ( m < i & f . i = b ) } ; ::_thesis: x in NAT then ex i being Element of NAT st ( x = i & m < i & f . i = b ) ; hence x in NAT ; ::_thesis: verum end; then reconsider N = { i where i is Element of NAT : ( m < i & f . i = b ) } as non empty Subset of NAT by A4, TARSKI:def_3; take I = min N; ::_thesis: ( f . I = b & m < I & ( for i being Element of NAT st m < i & f . i = b holds I <= i ) ) I in N by XXREAL_2:def_7; then ex II being Element of NAT st ( II = I & m < II & f . II = b ) ; hence ( f . I = b & m < I ) ; ::_thesis: for i being Element of NAT st m < i & f . i = b holds I <= i let i be Element of NAT ; ::_thesis: ( m < i & f . i = b implies I <= i ) assume that A5: m < i and A6: f . i = b ; ::_thesis: I <= i i in N by A5, A6; hence I <= i by XXREAL_2:def_7; ::_thesis: verum end; uniqueness for b1, b2 being Element of NAT st f . b1 = b & m < b1 & ( for i being Element of NAT st m < i & f . i = b holds b1 <= i ) & f . b2 = b & m < b2 & ( for i being Element of NAT st m < i & f . i = b holds b2 <= i ) holds b1 = b2 proof let IT1, IT2 be Element of NAT ; ::_thesis: ( f . IT1 = b & m < IT1 & ( for i being Element of NAT st m < i & f . i = b holds IT1 <= i ) & f . IT2 = b & m < IT2 & ( for i being Element of NAT st m < i & f . i = b holds IT2 <= i ) implies IT1 = IT2 ) assume that A7: f . IT1 = b and A8: m < IT1 and A9: for i being Element of NAT st m < i & f . i = b holds IT1 <= i and A10: f . IT2 = b and A11: m < IT2 and A12: for i being Element of NAT st m < i & f . i = b holds IT2 <= i ; ::_thesis: IT1 = IT2 assume A13: IT1 <> IT2 ; ::_thesis: contradiction percases ( IT1 < IT2 or IT1 > IT2 ) by A13, XXREAL_0:1; suppose IT1 < IT2 ; ::_thesis: contradiction hence contradiction by A7, A8, A12; ::_thesis: verum end; suppose IT1 > IT2 ; ::_thesis: contradiction hence contradiction by A9, A10, A11; ::_thesis: verum end; end; end; end; :: deftheorem Def12 defines mindex DICKSON:def_12_:_ for R being non empty 1-sorted for f being sequence of R for b being set for m being Element of NAT st ex j being Element of NAT st ( m < j & f . j = b ) holds for b5 being Element of NAT holds ( b5 = f mindex (b,m) iff ( f . b5 = b & m < b5 & ( for i being Element of NAT st m < i & f . i = b holds b5 <= i ) ) ); theorem Th29: :: DICKSON:29 for R being non empty RelStr st R is Dickson holds for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) proof let R be non empty RelStr ; ::_thesis: ( R is Dickson implies for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) ) assume A1: R is Dickson ; ::_thesis: for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) let f be sequence of R; ::_thesis: ex i, j being Element of NAT st ( i < j & f . i <= f . j ) set N = rng f; A2: dom f = NAT by FUNCT_2:def_1; consider B being set such that A3: B is_Dickson-basis_of rng f,R and A4: B is finite by A1, Def10; reconsider B = B as non empty finite set by A3, A4, Th26; defpred S1[ set ] means verum; deffunc H1( set ) -> Element of NAT = f mindex $1; set BI = { H1(b) where b is Element of B : S1[b] } ; A5: { H1(b) where b is Element of B : S1[b] } is finite from PRE_CIRC:sch_1(); set b = the Element of B; A6: f mindex the Element of B in { H1(b) where b is Element of B : S1[b] } ; now__::_thesis:_for_x_being_set_st_x_in__{__H1(b)_where_b_is_Element_of_B_:_S1[b]__}__holds_ x_in_NAT let x be set ; ::_thesis: ( x in { H1(b) where b is Element of B : S1[b] } implies x in NAT ) assume x in { H1(b) where b is Element of B : S1[b] } ; ::_thesis: x in NAT then ex b being Element of B st x = f mindex b ; hence x in NAT ; ::_thesis: verum end; then reconsider BI = { H1(b) where b is Element of B : S1[b] } as non empty finite Subset of NAT by A5, A6, TARSKI:def_3; reconsider mB = max BI as Element of NAT by ORDINAL1:def_12; set j = mB + 1; reconsider fj = f . (mB + 1) as Element of R ; fj in rng f by A2, FUNCT_1:3; then consider b being Element of R such that A7: b in B and A8: b <= fj by A3, Def9; A9: B c= rng f by A3, Def9; take i = f mindex b; ::_thesis: ex j being Element of NAT st ( i < j & f . i <= f . j ) take mB + 1 ; ::_thesis: ( i < mB + 1 & f . i <= f . (mB + 1) ) i in BI by A7; then i <= max BI by XXREAL_2:def_8; hence i < mB + 1 by NAT_1:13; ::_thesis: f . i <= f . (mB + 1) dom f = NAT by NORMSP_1:12; hence f . i <= f . (mB + 1) by A7, A8, A9, Def11; ::_thesis: verum end; theorem Th30: :: DICKSON:30 for R being RelStr for N being Subset of R for x being Element of (R \~) st R is quasi_ordered & x in N & ( the InternalRel of R -Seg x) /\ N c= Class ((EqRel R),x) holds x is_minimal_wrt N, the InternalRel of (R \~) proof let R be RelStr ; ::_thesis: for N being Subset of R for x being Element of (R \~) st R is quasi_ordered & x in N & ( the InternalRel of R -Seg x) /\ N c= Class ((EqRel R),x) holds x is_minimal_wrt N, the InternalRel of (R \~) let N be Subset of R; ::_thesis: for x being Element of (R \~) st R is quasi_ordered & x in N & ( the InternalRel of R -Seg x) /\ N c= Class ((EqRel R),x) holds x is_minimal_wrt N, the InternalRel of (R \~) let x be Element of (R \~); ::_thesis: ( R is quasi_ordered & x in N & ( the InternalRel of R -Seg x) /\ N c= Class ((EqRel R),x) implies x is_minimal_wrt N, the InternalRel of (R \~) ) assume that A1: R is quasi_ordered and A2: x in N and A3: ( the InternalRel of R -Seg x) /\ N c= Class ((EqRel R),x) ; ::_thesis: x is_minimal_wrt N, the InternalRel of (R \~) set IR = the InternalRel of R; set IR9 = the InternalRel of (R \~); now__::_thesis:_for_y_being_set_holds_ (_not_y_in_N_or_not_y_<>_x_or_not_[y,x]_in_the_InternalRel_of_(R_\~)_) assume ex y being set st ( y in N & y <> x & [y,x] in the InternalRel of (R \~) ) ; ::_thesis: contradiction then consider y being set such that A4: y in N and A5: y <> x and A6: [y,x] in the InternalRel of (R \~) ; A7: not [y,x] in the InternalRel of R ~ by A6, XBOOLE_0:def_5; y in the InternalRel of R -Seg x by A5, A6, WELLORD1:1; then y in ( the InternalRel of R -Seg x) /\ N by A4, XBOOLE_0:def_4; then [y,x] in EqRel R by A3, EQREL_1:19; then [y,x] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; hence contradiction by A7, XBOOLE_0:def_4; ::_thesis: verum end; hence x is_minimal_wrt N, the InternalRel of (R \~) by A2, WAYBEL_4:def_25; ::_thesis: verum end; theorem Th31: :: DICKSON:31 for R being non empty RelStr st R is quasi_ordered & ( for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) ) holds for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) proof let R be non empty RelStr ; ::_thesis: ( R is quasi_ordered & ( for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) ) implies for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) ) assume that A1: R is quasi_ordered and A2: for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) ; ::_thesis: for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) set IR = the InternalRel of R; set IR9 = the InternalRel of (R \~); A3: R is transitive by A1, Def3; let N be non empty Subset of R; ::_thesis: ( min-classes N is finite & not min-classes N is empty ) assume A4: ( not min-classes N is finite or min-classes N is empty ) ; ::_thesis: contradiction percases ( min-classes N is infinite or min-classes N is empty ) by A4; supposeA5: min-classes N is infinite ; ::_thesis: contradiction then reconsider MCN = min-classes N as infinite set ; consider f being Function of NAT,(min-classes N) such that A6: f is one-to-one by A5, Th3; deffunc H1( set ) -> Element of f . $1 = choose (f . $1); A7: now__::_thesis:_for_x_being_set_st_x_in_NAT_holds_ H1(x)_in_the_carrier_of_R let x be set ; ::_thesis: ( x in NAT implies H1(x) in the carrier of R ) assume x in NAT ; ::_thesis: H1(x) in the carrier of R then reconsider x9 = x as Element of NAT ; f . x9 is Element of min-classes N ; then A8: f . x in MCN ; then not f . x is empty by Th22; then choose (f . x9) in f . x ; hence H1(x) in the carrier of R by A8; ::_thesis: verum end; consider g being Function of NAT, the carrier of R such that A9: for x being set st x in NAT holds g . x = H1(x) from FUNCT_2:sch_2(A7); reconsider g = g as sequence of R ; consider i, j being Element of NAT such that A10: i < j and A11: g . i <= g . j by A2; reconsider gi = g . i, gj = g . j as Element of (R \~) ; A12: [gi,gj] in the InternalRel of R by A11, ORDERS_2:def_5; A13: f . i in MCN ; then A14: not f . i is empty by Th22; A15: f . j in MCN ; then A16: not f . j is empty by Th22; A17: g . j = choose (f . j) by A9; A18: g . i = choose (f . i) by A9; A19: gj is_minimal_wrt N, the InternalRel of (R \~) by A1, A15, A16, A17, Th19; gi is_minimal_wrt N, the InternalRel of (R \~) by A1, A13, A14, A18, Th19; then A20: gi in N by WAYBEL_4:def_25; A21: now__::_thesis:_[gi,gj]_in_the_InternalRel_of_R_~ percases ( gi = gj or gi <> gj ) ; suppose gi = gj ; ::_thesis: [gi,gj] in the InternalRel of R ~ hence [gi,gj] in the InternalRel of R ~ by A12, RELAT_1:def_7; ::_thesis: verum end; supposeA22: gi <> gj ; ::_thesis: [gi,gj] in the InternalRel of R ~ now__::_thesis:_[gi,gj]_in_the_InternalRel_of_R_~ assume not [gi,gj] in the InternalRel of R ~ ; ::_thesis: contradiction then [gi,gj] in the InternalRel of R \ ( the InternalRel of R ~) by A12, XBOOLE_0:def_5; hence contradiction by A19, A20, A22, WAYBEL_4:def_25; ::_thesis: verum end; hence [gi,gj] in the InternalRel of R ~ ; ::_thesis: verum end; end; end; [gi,gj] in the InternalRel of R by A11, ORDERS_2:def_5; then [gi,gj] in the InternalRel of R /\ ( the InternalRel of R ~) by A21, XBOOLE_0:def_4; then [gi,gj] in EqRel R by A1, Def4; then gi in Class ((EqRel R),gj) by EQREL_1:19; then A23: Class ((EqRel R),gj) = Class ((EqRel R),gi) by EQREL_1:23; consider mj being Element of (R \~) such that mj is_minimal_wrt N, the InternalRel of (R \~) and A24: f . j = (Class ((EqRel R),mj)) /\ N by A15, Def8; consider mi being Element of (R \~) such that mi is_minimal_wrt N, the InternalRel of (R \~) and A25: f . i = (Class ((EqRel R),mi)) /\ N by A13, Def8; gj in Class ((EqRel R),mj) by A16, A17, A24, XBOOLE_0:def_4; then A26: Class ((EqRel R),gj) = Class ((EqRel R),mj) by EQREL_1:23; gi in Class ((EqRel R),mi) by A14, A18, A25, XBOOLE_0:def_4; then f . i = f . j by A23, A24, A25, A26, EQREL_1:23; hence contradiction by A5, A6, A10, FUNCT_2:19; ::_thesis: verum end; supposeA27: min-classes N is empty ; ::_thesis: contradiction deffunc H1( set , set ) -> Element of (( the InternalRel of R -Seg $2) /\ N) \ (Class ((EqRel R),$2)) = choose ((( the InternalRel of R -Seg $2) /\ N) \ (Class ((EqRel R),$2))); consider f being Function such that A28: dom f = NAT and A29: f . 0 = choose N and A30: for n being Nat holds f . (n + 1) = H1(n,f . n) from NAT_1:sch_11(); defpred S1[ Nat] means f . $1 in N; A31: S1[ 0 ] by A29; A32: now__::_thesis:_for_i_being_Element_of_NAT_st_S1[i]_holds_ S1[i_+_1] let i be Element of NAT ; ::_thesis: ( S1[i] implies S1[i + 1] ) assume A33: S1[i] ; ::_thesis: S1[i + 1] reconsider fi = f . i as Element of (R \~) by A33; set IC = (( the InternalRel of R -Seg fi) /\ N) \ (Class ((EqRel R),fi)); A34: f . (i + 1) = choose ((( the InternalRel of R -Seg (f . i)) /\ N) \ (Class ((EqRel R),(f . i)))) by A30; now__::_thesis:_not_((_the_InternalRel_of_R_-Seg_fi)_/\_N)_\_(Class_((EqRel_R),fi))_is_empty assume (( the InternalRel of R -Seg fi) /\ N) \ (Class ((EqRel R),fi)) is empty ; ::_thesis: contradiction then ( the InternalRel of R -Seg fi) /\ N c= Class ((EqRel R),fi) by XBOOLE_1:37; hence contradiction by A1, A27, A33, Th21, Th30; ::_thesis: verum end; then f . (i + 1) in ( the InternalRel of R -Seg (f . i)) /\ N by A34, XBOOLE_0:def_5; hence S1[i + 1] by XBOOLE_0:def_4; ::_thesis: verum end; A35: for i being Element of NAT holds S1[i] from NAT_1:sch_1(A31, A32); now__::_thesis:_for_x_being_set_st_x_in_NAT_holds_ f_._x_in_the_carrier_of_R let x be set ; ::_thesis: ( x in NAT implies f . x in the carrier of R ) assume x in NAT ; ::_thesis: f . x in the carrier of R then f . x in N by A35; hence f . x in the carrier of R ; ::_thesis: verum end; then reconsider f = f as sequence of R by A28, FUNCT_2:3; A36: now__::_thesis:_for_i,_j_being_Element_of_NAT_holds_S2[j] let i be Element of NAT ; ::_thesis: for j being Element of NAT holds S2[j] defpred S2[ Element of NAT ] means ( i < $1 implies f . i >= f . $1 ); A37: S2[ 0 ] by NAT_1:2; A38: for j being Element of NAT st S2[j] holds S2[j + 1] proof let j be Element of NAT ; ::_thesis: ( S2[j] implies S2[j + 1] ) assume that A39: ( i < j implies f . i >= f . j ) and A40: i < j + 1 ; ::_thesis: f . i >= f . (j + 1) A41: i <= j by A40, NAT_1:13; reconsider fj = f . j, fj1 = f . (j + 1) as Element of (R \~) ; set IC = (( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj)); A42: fj in N by A35; A43: fj1 = choose ((( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj))) by A30; now__::_thesis:_not_((_the_InternalRel_of_R_-Seg_fj)_/\_N)_\_(Class_((EqRel_R),fj))_is_empty assume (( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj)) is empty ; ::_thesis: contradiction then ( the InternalRel of R -Seg fj) /\ N c= Class ((EqRel R),fj) by XBOOLE_1:37; hence contradiction by A1, A27, A42, Th21, Th30; ::_thesis: verum end; then fj1 in ( the InternalRel of R -Seg fj) /\ N by A43, XBOOLE_0:def_5; then fj1 in the InternalRel of R -Seg fj by XBOOLE_0:def_4; then A44: [fj1,fj] in the InternalRel of R by WELLORD1:1; then A45: f . j >= f . (j + 1) by ORDERS_2:def_5; percases ( i < j or i = j ) by A41, XXREAL_0:1; suppose i < j ; ::_thesis: f . i >= f . (j + 1) hence f . i >= f . (j + 1) by A3, A39, A45, ORDERS_2:3; ::_thesis: verum end; suppose i = j ; ::_thesis: f . i >= f . (j + 1) hence f . i >= f . (j + 1) by A44, ORDERS_2:def_5; ::_thesis: verum end; end; end; thus for j being Element of NAT holds S2[j] from NAT_1:sch_1(A37, A38); ::_thesis: verum end; now__::_thesis:_for_i,_j_being_Element_of_NAT_holds_S2[j] let i be Element of NAT ; ::_thesis: for j being Element of NAT holds S2[j] defpred S2[ Element of NAT ] means ( i < $1 implies not f . i <= f . $1 ); A46: S2[ 0 ] by NAT_1:2; A47: for j being Element of NAT st S2[j] holds S2[j + 1] proof let j be Element of NAT ; ::_thesis: ( S2[j] implies S2[j + 1] ) assume that ( i < j implies not f . i <= f . j ) and A48: i < j + 1 ; ::_thesis: not f . i <= f . (j + 1) A49: i <= j by A48, NAT_1:13; reconsider fj = f . j, fj1 = f . (j + 1) as Element of (R \~) ; A50: fj in N by A35; percases ( i < j or i = j ) by A49, XXREAL_0:1; supposeA51: i < j ; ::_thesis: not f . i <= f . (j + 1) assume A52: f . i <= f . (j + 1) ; ::_thesis: contradiction j < j + 1 by NAT_1:13; then A53: f . j >= f . (j + 1) by A36; f . i >= f . j by A36, A51; then f . j <= f . (j + 1) by A3, A52, ORDERS_2:3; then A54: fj1 in Class ((EqRel R),fj) by A1, A53, Th8; set IC = (( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj)); A55: fj1 = choose ((( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj))) by A30; now__::_thesis:_not_((_the_InternalRel_of_R_-Seg_fj)_/\_N)_\_(Class_((EqRel_R),fj))_is_empty assume (( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj)) is empty ; ::_thesis: contradiction then ( the InternalRel of R -Seg fj) /\ N c= Class ((EqRel R),fj) by XBOOLE_1:37; hence contradiction by A1, A27, A50, Th21, Th30; ::_thesis: verum end; hence contradiction by A54, A55, XBOOLE_0:def_5; ::_thesis: verum end; supposeA56: i = j ; ::_thesis: not f . i <= f . (j + 1) assume A57: f . i <= f . (j + 1) ; ::_thesis: contradiction j < j + 1 by NAT_1:13; then f . (j + 1) <= f . j by A36; then A58: fj1 in Class ((EqRel R),fj) by A1, A56, A57, Th8; set IC = (( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj)); A59: fj1 = choose ((( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj))) by A30; now__::_thesis:_not_((_the_InternalRel_of_R_-Seg_fj)_/\_N)_\_(Class_((EqRel_R),fj))_is_empty assume (( the InternalRel of R -Seg fj) /\ N) \ (Class ((EqRel R),fj)) is empty ; ::_thesis: contradiction then ( the InternalRel of R -Seg fj) /\ N c= Class ((EqRel R),fj) by XBOOLE_1:37; hence contradiction by A1, A27, A50, Th21, Th30; ::_thesis: verum end; hence contradiction by A58, A59, XBOOLE_0:def_5; ::_thesis: verum end; end; end; thus for j being Element of NAT holds S2[j] from NAT_1:sch_1(A46, A47); ::_thesis: verum end; hence contradiction by A2; ::_thesis: verum end; end; end; theorem Th32: :: DICKSON:32 for R being non empty RelStr st R is quasi_ordered & ( for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) ) holds R is Dickson proof let R be non empty RelStr ; ::_thesis: ( R is quasi_ordered & ( for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) ) implies R is Dickson ) assume that A1: R is quasi_ordered and A2: for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) ; ::_thesis: R is Dickson A3: R is transitive by A1, Def3; A4: R is reflexive by A1, Def3; set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); let N be Subset of the carrier of R; :: according to DICKSON:def_10 ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) percases ( N = {} the carrier of R or not N is empty ) ; supposeA5: N = {} the carrier of R ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) take B = {} ; ::_thesis: ( B is_Dickson-basis_of N,R & B is finite ) thus B is_Dickson-basis_of N,R by A5, Th25; ::_thesis: B is finite thus B is finite ; ::_thesis: verum end; suppose not N is empty ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) then reconsider N9 = N as non empty Subset of the carrier of R ; reconsider MCN9 = min-classes N9 as non empty finite Subset-Family of the carrier of R by A2; take B = { (choose x) where x is Element of MCN9 : verum } ; ::_thesis: ( B is_Dickson-basis_of N,R & B is finite ) thus B is_Dickson-basis_of N,R ::_thesis: B is finite proof now__::_thesis:_for_x_being_set_st_x_in_B_holds_ x_in_N let x be set ; ::_thesis: ( x in B implies x in N ) assume x in B ; ::_thesis: x in N then consider y being Element of MCN9 such that A6: x = choose y ; A7: ex z being Element of (R \~) st ( z is_minimal_wrt N, the InternalRel of (R \~) & y = (Class ((EqRel R),z)) /\ N ) by Def8; not y is empty by Th22; hence x in N by A6, A7, XBOOLE_0:def_4; ::_thesis: verum end; hence B c= N by TARSKI:def_3; :: according to DICKSON:def_9 ::_thesis: for a being Element of R st a in N holds ex b being Element of R st ( b in B & b <= a ) let a be Element of R; ::_thesis: ( a in N implies ex b being Element of R st ( b in B & b <= a ) ) assume A8: a in N ; ::_thesis: ex b being Element of R st ( b in B & b <= a ) defpred S1[ Element of R] means $1 <= a; set NN9 = { d where d is Element of N9 : S1[d] } ; A9: { d where d is Element of N9 : S1[d] } is Subset of N9 from DOMAIN_1:sch_7(); a <= a by A4, ORDERS_2:1; then a in { d where d is Element of N9 : S1[d] } by A8; then reconsider NN9 = { d where d is Element of N9 : S1[d] } as non empty Subset of the carrier of R by A9, XBOOLE_1:1; set m = the Element of min-classes NN9; A10: not min-classes NN9 is empty by A2; then reconsider m9 = the Element of min-classes NN9 as non empty set by Th22; set c = the Element of m9; consider y being Element of (R \~) such that y is_minimal_wrt NN9, the InternalRel of (R \~) and A11: m9 = (Class ((EqRel R),y)) /\ NN9 by A10, Def8; A12: the Element of m9 in Class ((EqRel R),y) by A11, XBOOLE_0:def_4; A13: the Element of m9 in NN9 by A11, XBOOLE_0:def_4; reconsider c = the Element of m9 as Element of (R \~) by A12; reconsider c9 = c as Element of R ; A14: ex d being Element of N9 st ( c = d & d <= a ) by A13; A15: c is_minimal_wrt NN9, the InternalRel of (R \~) by A1, A10, Th19; now__::_thesis:_c_is_minimal_wrt_N,_the_InternalRel_of_(R_\~) assume not c is_minimal_wrt N, the InternalRel of (R \~) ; ::_thesis: contradiction then consider w being set such that A16: w in N and A17: w <> c and A18: [w,c] in the InternalRel of (R \~) by A14, WAYBEL_4:def_25; reconsider w9 = w as Element of R by A18, ZFMISC_1:87; w9 <= c9 by A18, ORDERS_2:def_5; then w9 <= a by A3, A14, ORDERS_2:3; then w9 in NN9 by A16; hence contradiction by A15, A17, A18, WAYBEL_4:def_25; ::_thesis: verum end; then A19: (Class ((EqRel R),c)) /\ N in min-classes N by Def8; then A20: not (Class ((EqRel R),c)) /\ N is empty by Th22; set t = choose ((Class ((EqRel R),c)) /\ N); choose ((Class ((EqRel R),c)) /\ N) in N by A20, XBOOLE_0:def_4; then reconsider t = choose ((Class ((EqRel R),c)) /\ N) as Element of R ; take t ; ::_thesis: ( t in B & t <= a ) thus t in B by A19; ::_thesis: t <= a t in Class ((EqRel R),c) by A20, XBOOLE_0:def_4; then [t,c] in EqRel R by EQREL_1:19; then [t,c] in the InternalRel of R /\ ( the InternalRel of R ~) by A1, Def4; then [t,c] in the InternalRel of R by XBOOLE_0:def_4; then t <= c9 by ORDERS_2:def_5; hence t <= a by A3, A14, ORDERS_2:3; ::_thesis: verum end; deffunc H1( set ) -> Element of $1 = choose $1; defpred S1[ set ] means verum; { H1(x) where x is Element of MCN9 : S1[x] } is finite from PRE_CIRC:sch_1(); hence B is finite ; ::_thesis: verum end; end; end; theorem Th33: :: DICKSON:33 for R being non empty RelStr st R is quasi_ordered & R is Dickson holds R \~ is well_founded proof let R be non empty RelStr ; ::_thesis: ( R is quasi_ordered & R is Dickson implies R \~ is well_founded ) assume that A1: R is quasi_ordered and A2: R is Dickson ; ::_thesis: R \~ is well_founded A3: for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) by A2, Th29; now__::_thesis:_for_N_being_Subset_of_R_st_N_<>_{}_holds_ ex_x_being_set_st_x_in_min-classes_N let N be Subset of R; ::_thesis: ( N <> {} implies ex x being set st x in min-classes N ) assume N <> {} ; ::_thesis: ex x being set st x in min-classes N then not min-classes N is empty by A1, A3, Th31; hence ex x being set st x in min-classes N by XBOOLE_0:def_1; ::_thesis: verum end; hence R \~ is well_founded by Th20; ::_thesis: verum end; theorem :: DICKSON:34 for R being non empty Poset for N being non empty Subset of R st R is Dickson holds ex B being set st ( B is_Dickson-basis_of N,R & ( for C being set st C is_Dickson-basis_of N,R holds B c= C ) ) proof let R be non empty Poset; ::_thesis: for N being non empty Subset of R st R is Dickson holds ex B being set st ( B is_Dickson-basis_of N,R & ( for C being set st C is_Dickson-basis_of N,R holds B c= C ) ) let N be non empty Subset of R; ::_thesis: ( R is Dickson implies ex B being set st ( B is_Dickson-basis_of N,R & ( for C being set st C is_Dickson-basis_of N,R holds B c= C ) ) ) assume A1: R is Dickson ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & ( for C being set st C is_Dickson-basis_of N,R holds B c= C ) ) set IR = the InternalRel of R; set CR = the carrier of R; set IR9 = the InternalRel of (R \~); set B = { b where b is Element of (R \~) : b is_minimal_wrt N, the InternalRel of (R \~) } ; A2: R is quasi_ordered by Def3; for f being sequence of R ex i, j being Element of NAT st ( i < j & f . i <= f . j ) by A1, Th29; then not min-classes N is empty by A2, Th31; then consider x being set such that A3: x in min-classes N by XBOOLE_0:def_1; consider y being Element of (R \~) such that A4: y is_minimal_wrt N, the InternalRel of (R \~) and x = (Class ((EqRel R),y)) /\ N by A3, Def8; y in { b where b is Element of (R \~) : b is_minimal_wrt N, the InternalRel of (R \~) } by A4; then reconsider B = { b where b is Element of (R \~) : b is_minimal_wrt N, the InternalRel of (R \~) } as non empty set ; take B ; ::_thesis: ( B is_Dickson-basis_of N,R & ( for C being set st C is_Dickson-basis_of N,R holds B c= C ) ) A5: now__::_thesis:_for_x_being_set_st_x_in_B_holds_ x_in_N let x be set ; ::_thesis: ( x in B implies x in N ) assume x in B ; ::_thesis: x in N then ex b being Element of (R \~) st ( x = b & b is_minimal_wrt N, the InternalRel of (R \~) ) ; hence x in N by WAYBEL_4:def_25; ::_thesis: verum end; then A6: B c= N by TARSKI:def_3; thus B is_Dickson-basis_of N,R ::_thesis: for C being set st C is_Dickson-basis_of N,R holds B c= C proof thus B c= N by A5, TARSKI:def_3; :: according to DICKSON:def_9 ::_thesis: for a being Element of R st a in N holds ex b being Element of R st ( b in B & b <= a ) let a be Element of R; ::_thesis: ( a in N implies ex b being Element of R st ( b in B & b <= a ) ) assume A7: a in N ; ::_thesis: ex b being Element of R st ( b in B & b <= a ) reconsider a9 = a as Element of (R \~) ; now__::_thesis:_ex_b_being_Element_of_R_st_ (_b_in_B_&_b_<=_a_) assume A8: for b being Element of R holds ( not b in B or not b <= a ) ; ::_thesis: contradiction percases ( ( the InternalRel of (R \~) -Seg a) /\ N = {} or ( the InternalRel of (R \~) -Seg a) /\ N <> {} ) ; suppose ( the InternalRel of (R \~) -Seg a) /\ N = {} ; ::_thesis: contradiction then the InternalRel of (R \~) -Seg a misses N by XBOOLE_0:def_7; then a9 is_minimal_wrt N, the InternalRel of (R \~) by A7, Th6; then a in B ; hence contradiction by A8; ::_thesis: verum end; supposeA9: ( the InternalRel of (R \~) -Seg a) /\ N <> {} ; ::_thesis: contradiction R \~ is well_founded by A1, A2, Th33; then the InternalRel of (R \~) is_well_founded_in the carrier of R by WELLFND1:def_2; then consider z being set such that A10: z in ( the InternalRel of (R \~) -Seg a) /\ N and A11: the InternalRel of (R \~) -Seg z misses ( the InternalRel of (R \~) -Seg a) /\ N by A9, WELLORD1:def_3; A12: z in the InternalRel of (R \~) -Seg a by A10, XBOOLE_0:def_4; then [z,a] in the InternalRel of (R \~) by WELLORD1:1; then z in dom the InternalRel of (R \~) by XTUPLE_0:def_12; then reconsider z = z as Element of (R \~) ; reconsider z9 = z as Element of R ; z is_minimal_wrt ( the InternalRel of (R \~) -Seg a9) /\ N, the InternalRel of (R \~) by A10, A11, Th6; then z is_minimal_wrt N, the InternalRel of (R \~) by Th7; then A13: z in B ; [z,a] in the InternalRel of R \ ( the InternalRel of R ~) by A12, WELLORD1:1; then z9 <= a by ORDERS_2:def_5; hence contradiction by A8, A13; ::_thesis: verum end; end; end; hence ex b being Element of R st ( b in B & b <= a ) ; ::_thesis: verum end; let C be set ; ::_thesis: ( C is_Dickson-basis_of N,R implies B c= C ) assume A14: C is_Dickson-basis_of N,R ; ::_thesis: B c= C A15: C c= N by A14, Def9; now__::_thesis:_for_b_being_set_st_b_in_B_holds_ b_in_C let b be set ; ::_thesis: ( b in B implies b in C ) assume A16: b in B ; ::_thesis: b in C b in N by A5, A16; then reconsider b9 = b as Element of R ; consider c being Element of R such that A17: c in C and A18: c <= b9 by A6, A14, A16, Def9; A19: ex b99 being Element of (R \~) st ( b99 = b & b99 is_minimal_wrt N, the InternalRel of (R \~) ) by A16; A20: [c,b] in the InternalRel of R by A18, ORDERS_2:def_5; A21: the InternalRel of R is_antisymmetric_in the carrier of R by ORDERS_2:def_4; [b,c] in the InternalRel of R by A15, A17, A19, A20, Th17; hence b in C by A17, A18, A20, A21, RELAT_2:def_4; ::_thesis: verum end; hence B c= C by TARSKI:def_3; ::_thesis: verum end; definition let R be non empty RelStr ; let N be Subset of R; assume B1: R is Dickson ; func Dickson-bases (N,R) -> non empty Subset-Family of R means :Def13: :: DICKSON:def 13 for B being set holds ( B in it iff B is_Dickson-basis_of N,R ); existence ex b1 being non empty Subset-Family of R st for B being set holds ( B in b1 iff B is_Dickson-basis_of N,R ) proof set BB = { b where b is Subset of N : b is_Dickson-basis_of N,R } ; set CR = the carrier of R; consider bp being set such that A1: bp is_Dickson-basis_of N,R and bp is finite by B1, Def10; bp c= N by A1, Def9; then A2: bp in { b where b is Subset of N : b is_Dickson-basis_of N,R } by A1; now__::_thesis:_for_x_being_set_st_x_in__{__b_where_b_is_Subset_of_N_:_b_is_Dickson-basis_of_N,R__}__holds_ x_in_bool_the_carrier_of_R let x be set ; ::_thesis: ( x in { b where b is Subset of N : b is_Dickson-basis_of N,R } implies x in bool the carrier of R ) assume x in { b where b is Subset of N : b is_Dickson-basis_of N,R } ; ::_thesis: x in bool the carrier of R then consider b being Subset of N such that A3: x = b and b is_Dickson-basis_of N,R ; b c= the carrier of R by XBOOLE_1:1; hence x in bool the carrier of R by A3; ::_thesis: verum end; then reconsider BB = { b where b is Subset of N : b is_Dickson-basis_of N,R } as non empty Subset-Family of the carrier of R by A2, TARSKI:def_3; take BB ; ::_thesis: for B being set holds ( B in BB iff B is_Dickson-basis_of N,R ) let B be set ; ::_thesis: ( B in BB iff B is_Dickson-basis_of N,R ) hereby ::_thesis: ( B is_Dickson-basis_of N,R implies B in BB ) assume B in BB ; ::_thesis: B is_Dickson-basis_of N,R then ex b being Subset of N st ( b = B & b is_Dickson-basis_of N,R ) ; hence B is_Dickson-basis_of N,R ; ::_thesis: verum end; assume A4: B is_Dickson-basis_of N,R ; ::_thesis: B in BB then B c= N by Def9; hence B in BB by A4; ::_thesis: verum end; uniqueness for b1, b2 being non empty Subset-Family of R st ( for B being set holds ( B in b1 iff B is_Dickson-basis_of N,R ) ) & ( for B being set holds ( B in b2 iff B is_Dickson-basis_of N,R ) ) holds b1 = b2 proof let IT1, IT2 be non empty Subset-Family of R; ::_thesis: ( ( for B being set holds ( B in IT1 iff B is_Dickson-basis_of N,R ) ) & ( for B being set holds ( B in IT2 iff B is_Dickson-basis_of N,R ) ) implies IT1 = IT2 ) assume that A5: for B being set holds ( B in IT1 iff B is_Dickson-basis_of N,R ) and A6: for B being set holds ( B in IT2 iff B is_Dickson-basis_of N,R ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_x_being_set_holds_ (_(_x_in_IT1_implies_x_in_IT2_)_&_(_x_in_IT2_implies_x_in_IT1_)_) let x be set ; ::_thesis: ( ( x in IT1 implies x in IT2 ) & ( x in IT2 implies x in IT1 ) ) hereby ::_thesis: ( x in IT2 implies x in IT1 ) assume x in IT1 ; ::_thesis: x in IT2 then x is_Dickson-basis_of N,R by A5; hence x in IT2 by A6; ::_thesis: verum end; assume x in IT2 ; ::_thesis: x in IT1 then x is_Dickson-basis_of N,R by A6; hence x in IT1 by A5; ::_thesis: verum end; hence IT1 = IT2 by TARSKI:1; ::_thesis: verum end; end; :: deftheorem Def13 defines Dickson-bases DICKSON:def_13_:_ for R being non empty RelStr for N being Subset of R st R is Dickson holds for b3 being non empty Subset-Family of R holds ( b3 = Dickson-bases (N,R) iff for B being set holds ( B in b3 iff B is_Dickson-basis_of N,R ) ); theorem Th35: :: DICKSON:35 for R being non empty RelStr for s being sequence of R st R is Dickson holds ex t being sequence of R st ( t is subsequence of s & t is weakly-ascending ) proof let R be non empty RelStr ; ::_thesis: for s being sequence of R st R is Dickson holds ex t being sequence of R st ( t is subsequence of s & t is weakly-ascending ) let s be sequence of R; ::_thesis: ( R is Dickson implies ex t being sequence of R st ( t is subsequence of s & t is weakly-ascending ) ) assume A1: R is Dickson ; ::_thesis: ex t being sequence of R st ( t is subsequence of s & t is weakly-ascending ) set CR = the carrier of R; deffunc H1( Element of rng s, Element of NAT ) -> set = { n where n is Element of NAT : ( $1 <= s . n & $2 < n ) } ; deffunc H2( Element of rng s) -> set = { n where n is Element of NAT : $1 <= s . n } ; defpred S1[ set , Element of NAT , set ] means ex N being Subset of the carrier of R ex B being non empty Subset of the carrier of R st ( N = { (s . w) where w is Element of NAT : ( s . $2 <= s . w & $2 < w ) } & { w where w is Element of NAT : ( s . $2 <= s . w & $2 < w ) } is infinite & B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } & $3 = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,$2) is infinite ) } ),$2) ); defpred S2[ set , Element of NAT , set ] means ( { w where w is Element of NAT : ( s . $2 <= s . w & $2 < w ) } is infinite implies S1[$1,$2,$3] ); A2: for n, x being Element of NAT ex y being Element of NAT st S2[n,x,y] proof let n, x be Element of NAT ; ::_thesis: ex y being Element of NAT st S2[n,x,y] set N = { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } ; now__::_thesis:_for_y_being_set_st_y_in__{__(s_._w)_where_w_is_Element_of_NAT_:_(_s_._x_<=_s_._w_&_x_<_w_)__}__holds_ y_in_the_carrier_of_R let y be set ; ::_thesis: ( y in { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } implies y in the carrier of R ) assume y in { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } ; ::_thesis: y in the carrier of R then ex w being Element of NAT st ( y = s . w & s . x <= s . w & x < w ) ; hence y in the carrier of R ; ::_thesis: verum end; then reconsider N = { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } as Subset of the carrier of R by TARSKI:def_3; set W = { w where w is Element of NAT : ( s . x <= s . w & x < w ) } ; percases ( N is empty or not N is empty ) ; supposeA3: N is empty ; ::_thesis: ex y being Element of NAT st S2[n,x,y] take 1 ; ::_thesis: S2[n,x,1] assume { w where w is Element of NAT : ( s . x <= s . w & x < w ) } is infinite ; ::_thesis: S1[n,x,1] then consider ww being set such that A4: ww in { w where w is Element of NAT : ( s . x <= s . w & x < w ) } by XBOOLE_0:def_1; consider www being Element of NAT such that www = ww and A5: s . x <= s . www and A6: x < www by A4; s . www in N by A5, A6; hence S1[n,x,1] by A3; ::_thesis: verum end; supposeA7: not N is empty ; ::_thesis: ex y being Element of NAT st S2[n,x,y] set BBX = { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; set B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; consider BD1 being set such that A8: BD1 is_Dickson-basis_of N,R and A9: BD1 is finite by A1, Def10; BD1 in Dickson-bases (N,R) by A1, A8, Def13; then BD1 in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A9; then choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; then ex BBB being Element of Dickson-bases (N,R) st ( choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } = BBB & BBB is finite ) ; then A10: choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } is_Dickson-basis_of N,R by A1, Def13; then choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } c= N by Def9; then reconsider B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } as non empty Subset of the carrier of R by A7, A10, Th26, XBOOLE_1:1; set y = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x); take s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) ; ::_thesis: S2[n,x,s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x)] set W = { w where w is Element of NAT : ( s . x <= s . w & x < w ) } ; assume A11: { w where w is Element of NAT : ( s . x <= s . w & x < w ) } is infinite ; ::_thesis: S1[n,x,s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x)] take N ; ::_thesis: ex B being non empty Subset of the carrier of R st ( N = { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } & { w where w is Element of NAT : ( s . x <= s . w & x < w ) } is infinite & B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } & s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) ) reconsider B = B as non empty Subset of the carrier of R ; take B ; ::_thesis: ( N = { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } & { w where w is Element of NAT : ( s . x <= s . w & x < w ) } is infinite & B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } & s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) ) thus N = { (s . w) where w is Element of NAT : ( s . x <= s . w & x < w ) } ; ::_thesis: ( { w where w is Element of NAT : ( s . x <= s . w & x < w ) } is infinite & B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } & s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) ) thus { w where w is Element of NAT : ( s . x <= s . w & x < w ) } is infinite by A11; ::_thesis: ( B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } & s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) ) thus B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; ::_thesis: s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) thus s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,x) is infinite ) } ),x) ; ::_thesis: verum end; end; end; consider B being set such that A12: B is_Dickson-basis_of rng s,R and A13: B is finite by A1, Def10; reconsider B = B as non empty set by A12, Th26; set BALL = { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } ; set b1 = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } ; consider f being Function of NAT,NAT such that A14: f . 0 = s mindex (choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } ) and A15: for n being Element of NAT holds S2[n,f . n,f . (n + 1)] from RECDEF_1:sch_2(A2); A16: dom f = NAT by FUNCT_2:def_1; A17: rng f c= NAT ; now__::_thesis:_ex_b_being_Element_of_rng_s_st_ (_b_in_B_&_not_H2(b)_is_finite_) A18: B is finite by A13; assume A19: for b being Element of rng s st b in B holds H2(b) is finite ; ::_thesis: contradiction set Ball = { H2(b) where b is Element of rng s : b in B } ; A20: { H2(b) where b is Element of rng s : b in B } is finite from FRAENKEL:sch_21(A18); now__::_thesis:_for_X_being_set_st_X_in__{__H2(b)_where_b_is_Element_of_rng_s_:_b_in_B__}__holds_ X_is_finite let X be set ; ::_thesis: ( X in { H2(b) where b is Element of rng s : b in B } implies X is finite ) assume X in { H2(b) where b is Element of rng s : b in B } ; ::_thesis: X is finite then ex b being Element of rng s st ( X = H2(b) & b in B ) ; hence X is finite by A19; ::_thesis: verum end; then A21: union { H2(b) where b is Element of rng s : b in B } is finite by A20, FINSET_1:7; now__::_thesis:_for_x_being_set_st_x_in_NAT_holds_ x_in_union__{__H2(b)_where_b_is_Element_of_rng_s_:_b_in_B__}_ let x be set ; ::_thesis: ( x in NAT implies x in union { H2(b) where b is Element of rng s : b in B } ) assume x in NAT ; ::_thesis: x in union { H2(b) where b is Element of rng s : b in B } then reconsider x9 = x as Element of NAT ; dom s = NAT by FUNCT_2:def_1; then x9 in dom s ; then A22: s . x in rng s by FUNCT_1:3; then reconsider sx = s . x as Element of R ; consider b being Element of R such that A23: b in B and A24: b <= sx by A12, A22, Def9; B c= rng s by A12, Def9; then reconsider b = b as Element of rng s by A23; A25: x9 in H2(b) by A24; H2(b) in { H2(b) where b is Element of rng s : b in B } by A23; hence x in union { H2(b) where b is Element of rng s : b in B } by A25, TARSKI:def_4; ::_thesis: verum end; then NAT c= union { H2(b) where b is Element of rng s : b in B } by TARSKI:def_3; hence contradiction by A21; ::_thesis: verum end; then consider tb being Element of rng s such that A26: tb in B and A27: H2(tb) is infinite ; A28: tb in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } by A26, A27; then choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } ; then A29: ex bt being Element of B st ( choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } = bt & ex b9 being Element of rng s st ( b9 = bt & H2(b9) is infinite ) ) ; dom s = NAT by NORMSP_1:12; then A30: s . (f . 0) = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } by A14, A29, Def11; choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } by A28; then A31: ex eb being Element of B st ( choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H2(b9) is infinite ) } = eb & ex eb9 being Element of rng s st ( eb9 = eb & H2(eb9) is infinite ) ) ; deffunc H3( set ) -> set = $1; defpred S3[ Element of NAT ] means s . (f . 0) <= s . $1; set W1 = { w where w is Element of NAT : s . (f . 0) <= s . w } ; set W2 = { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } ; set W3 = { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } ; A32: { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } is finite from FINSEQ_1:sch_6(); now__::_thesis:_for_x_being_set_holds_ (_(_x_in__{__w_where_w_is_Element_of_NAT_:_s_._(f_._0)_<=_s_._w__}__implies_x_in__{__w_where_w_is_Element_of_NAT_:_(_s_._(f_._0)_<=_s_._w_&_f_._0_<_w_)__}__\/__{__H3(w)_where_w_is_Element_of_NAT_:_(_0_<=_w_&_w_<=_f_._0_&_S3[w]_)__}__)_&_(_x_in__{__w_where_w_is_Element_of_NAT_:_(_s_._(f_._0)_<=_s_._w_&_f_._0_<_w_)__}__\/__{__H3(w)_where_w_is_Element_of_NAT_:_(_0_<=_w_&_w_<=_f_._0_&_S3[w]_)__}__implies_x_in__{__w_where_w_is_Element_of_NAT_:_s_._(f_._0)_<=_s_._w__}__)_) let x be set ; ::_thesis: ( ( x in { w where w is Element of NAT : s . (f . 0) <= s . w } implies x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } \/ { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } ) & ( x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } \/ { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } implies b1 in { b2 where w is Element of NAT : s . (f . 0) <= s . b2 } ) ) hereby ::_thesis: ( x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } \/ { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } implies b1 in { b2 where w is Element of NAT : s . (f . 0) <= s . b2 } ) assume x in { w where w is Element of NAT : s . (f . 0) <= s . w } ; ::_thesis: x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } \/ { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } then consider w being Element of NAT such that A33: x = w and A34: s . (f . 0) <= s . w ; A35: 0 <= w by NAT_1:2; ( w <= f . 0 or w > f . 0 ) ; then ( x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } or x in { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } ) by A33, A34, A35; hence x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } \/ { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } by XBOOLE_0:def_3; ::_thesis: verum end; assume A36: x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } \/ { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } ; ::_thesis: b1 in { b2 where w is Element of NAT : s . (f . 0) <= s . b2 } percases ( x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } or x in { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } ) by A36, XBOOLE_0:def_3; suppose x in { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } ; ::_thesis: b1 in { b2 where w is Element of NAT : s . (f . 0) <= s . b2 } then ex w being Element of NAT st ( x = w & s . (f . 0) <= s . w & f . 0 < w ) ; hence x in { w where w is Element of NAT : s . (f . 0) <= s . w } ; ::_thesis: verum end; suppose x in { H3(w) where w is Element of NAT : ( 0 <= w & w <= f . 0 & S3[w] ) } ; ::_thesis: b1 in { b2 where w is Element of NAT : s . (f . 0) <= s . b2 } then ex w being Element of NAT st ( x = w & 0 <= w & w <= f . 0 & s . (f . 0) <= s . w ) ; hence x in { w where w is Element of NAT : s . (f . 0) <= s . w } ; ::_thesis: verum end; end; end; then A37: { w where w is Element of NAT : ( s . (f . 0) <= s . w & f . 0 < w ) } is infinite by A30, A31, A32, TARSKI:1; defpred S4[ Element of NAT ] means S1[$1,f . $1,f . ($1 + 1)]; A38: S4[ 0 ] by A15, A37; A39: now__::_thesis:_for_k_being_Element_of_NAT_st_S4[k]_holds_ S4[k_+_1] let k be Element of NAT ; ::_thesis: ( S4[k] implies S4[k + 1] ) assume S4[k] ; ::_thesis: S4[k + 1] then consider N being Subset of the carrier of R, B being non empty Subset of the carrier of R such that A40: N = { (s . w) where w is Element of NAT : ( s . (f . k) <= s . w & f . k < w ) } and A41: { w where w is Element of NAT : ( s . (f . k) <= s . w & f . k < w ) } is infinite and A42: B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } and A43: f . (k + 1) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } ),(f . k)) ; set BALL = { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } ; set BBX = { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; set iN = { w where w is Element of NAT : ( s . (f . k) <= s . w & f . k < w ) } ; consider BD being set such that A44: BD is_Dickson-basis_of N,R and A45: BD is finite by A1, Def10; BD in Dickson-bases (N,R) by A1, A44, Def13; then BD in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A45; then B in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A42; then A46: ex BB being Element of Dickson-bases (N,R) st ( B = BB & BB is finite ) ; then A47: B is_Dickson-basis_of N,R by A1, Def13; now__::_thesis:_ex_b_being_Element_of_rng_s_st_ (_b_in_B_&_not_H1(b,f_._k)_is_finite_) deffunc H4( Element of rng s) -> set = H1($1,f . k); A48: B is finite by A46; assume A49: for b being Element of rng s st b in B holds H1(b,f . k) is finite ; ::_thesis: contradiction set Ball = { H4(b) where b is Element of rng s : b in B } ; A50: { H4(b) where b is Element of rng s : b in B } is finite from FRAENKEL:sch_21(A48); now__::_thesis:_for_X_being_set_st_X_in__{__H4(b)_where_b_is_Element_of_rng_s_:_b_in_B__}__holds_ X_is_finite let X be set ; ::_thesis: ( X in { H4(b) where b is Element of rng s : b in B } implies X is finite ) assume X in { H4(b) where b is Element of rng s : b in B } ; ::_thesis: X is finite then ex b being Element of rng s st ( X = H1(b,f . k) & b in B ) ; hence X is finite by A49; ::_thesis: verum end; then A51: union { H4(b) where b is Element of rng s : b in B } is finite by A50, FINSET_1:7; { w where w is Element of NAT : ( s . (f . k) <= s . w & f . k < w ) } c= union { H4(b) where b is Element of rng s : b in B } proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { w where w is Element of NAT : ( s . (f . k) <= s . w & f . k < w ) } or x in union { H4(b) where b is Element of rng s : b in B } ) assume x in { w where w is Element of NAT : ( s . (f . k) <= s . w & f . k < w ) } ; ::_thesis: x in union { H4(b) where b is Element of rng s : b in B } then consider w being Element of NAT such that A52: x = w and A53: s . (f . k) <= s . w and A54: f . k < w ; A55: s . w in N by A40, A53, A54; reconsider sw = s . w as Element of R ; consider b being Element of R such that A56: b in B and A57: b <= sw by A47, A55, Def9; A58: B c= N by A47, Def9; N c= rng s proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in N or x in rng s ) assume x in N ; ::_thesis: x in rng s then A59: ex u being Element of NAT st ( x = s . u & s . (f . k) <= s . u & f . k < u ) by A40; dom s = NAT by FUNCT_2:def_1; hence x in rng s by A59, FUNCT_1:3; ::_thesis: verum end; then B c= rng s by A58, XBOOLE_1:1; then reconsider b = b as Element of rng s by A56; A60: w in H1(b,f . k) by A54, A57; H1(b,f . k) in { H4(b) where b is Element of rng s : b in B } by A56; hence x in union { H4(b) where b is Element of rng s : b in B } by A52, A60, TARSKI:def_4; ::_thesis: verum end; hence contradiction by A41, A51; ::_thesis: verum end; then consider b being Element of rng s such that A61: b in B and A62: H1(b,f . k) is infinite ; b in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } by A61, A62; then choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } ; then consider b being Element of B such that A63: b = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } and A64: ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) ; A65: b in B ; B c= N by A47, Def9; then b in N by A65; then A66: ex w being Element of NAT st ( b = s . w & s . (f . k) <= s . w & f . k < w ) by A40; then A67: s . (f . (k + 1)) = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . k) is infinite ) } by A43, A63, Def12; A68: f . k < f . (k + 1) by A43, A63, A66, Def12; deffunc H4( set ) -> set = $1; defpred S5[ Element of NAT ] means s . (f . (k + 1)) <= s . $1; set W1 = { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . k < w1 ) } ; set W2 = { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } ; set W3 = { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } ; A69: { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } is finite from DICKSON:sch_1(); now__::_thesis:_for_x_being_set_holds_ (_(_x_in__{__w1_where_w1_is_Element_of_NAT_:_(_s_._(f_._(k_+_1))_<=_s_._w1_&_f_._k_<_w1_)__}__implies_x_in__{__w1_where_w1_is_Element_of_NAT_:_(_s_._(f_._(k_+_1))_<=_s_._w1_&_f_._(k_+_1)_<_w1_)__}__\/__{__H4(w1)_where_w1_is_Element_of_NAT_:_(_f_._k_<_w1_&_w1_<=_f_._(k_+_1)_&_S5[w1]_)__}__)_&_(_x_in__{__w1_where_w1_is_Element_of_NAT_:_(_s_._(f_._(k_+_1))_<=_s_._w1_&_f_._(k_+_1)_<_w1_)__}__\/__{__H4(w1)_where_w1_is_Element_of_NAT_:_(_f_._k_<_w1_&_w1_<=_f_._(k_+_1)_&_S5[w1]_)__}__implies_x_in__{__w1_where_w1_is_Element_of_NAT_:_(_s_._(f_._(k_+_1))_<=_s_._w1_&_f_._k_<_w1_)__}__)_) let x be set ; ::_thesis: ( ( x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . k < w1 ) } implies x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } \/ { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } ) & ( x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } \/ { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } implies b1 in { b2 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . b2 & f . k < b2 ) } ) ) hereby ::_thesis: ( x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } \/ { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } implies b1 in { b2 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . b2 & f . k < b2 ) } ) assume x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . k < w1 ) } ; ::_thesis: x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } \/ { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } then consider w being Element of NAT such that A70: x = w and A71: s . (f . (k + 1)) <= s . w and A72: f . k < w ; ( w <= f . (k + 1) or w > f . (k + 1) ) ; then ( x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } or x in { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } ) by A70, A71, A72; hence x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } \/ { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } by XBOOLE_0:def_3; ::_thesis: verum end; assume A73: x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } \/ { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } ; ::_thesis: b1 in { b2 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . b2 & f . k < b2 ) } percases ( x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } or x in { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } ) by A73, XBOOLE_0:def_3; suppose x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } ; ::_thesis: b1 in { b2 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . b2 & f . k < b2 ) } then consider w being Element of NAT such that A74: x = w and A75: s . (f . (k + 1)) <= s . w and A76: f . (k + 1) < w ; f . k < w by A68, A76, XXREAL_0:2; hence x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . k < w1 ) } by A74, A75; ::_thesis: verum end; suppose x in { H4(w1) where w1 is Element of NAT : ( f . k < w1 & w1 <= f . (k + 1) & S5[w1] ) } ; ::_thesis: b1 in { b2 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . b2 & f . k < b2 ) } then ex w being Element of NAT st ( x = w & f . k < w & w <= f . (k + 1) & s . (f . (k + 1)) <= s . w ) ; hence x in { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . k < w1 ) } ; ::_thesis: verum end; end; end; then { w1 where w1 is Element of NAT : ( s . (f . (k + 1)) <= s . w1 & f . (k + 1) < w1 ) } is infinite by A63, A64, A67, A69, TARSKI:1; hence S4[k + 1] by A15; ::_thesis: verum end; A77: for n being Element of NAT holds S4[n] from NAT_1:sch_1(A38, A39); set t = s * f; take s * f ; ::_thesis: ( s * f is subsequence of s & s * f is weakly-ascending ) reconsider f = f as Function of NAT,REAL by FUNCT_2:7; now__::_thesis:_(_f_is_increasing_&_(_for_n_being_Element_of_NAT_holds_f_._n_is_Element_of_NAT_)_) now__::_thesis:_for_n_being_Element_of_NAT_holds_f_._n_<_f_._(n_+_1) let n be Element of NAT ; ::_thesis: f . n < f . (n + 1) f . n in rng f by A16, FUNCT_1:def_3; then reconsider fn = f . n as Element of NAT by A17; consider N being Subset of the carrier of R, B being non empty Subset of the carrier of R such that A78: N = { (s . w) where w is Element of NAT : ( s . fn <= s . w & fn < w ) } and A79: { w where w is Element of NAT : ( s . fn <= s . w & fn < w ) } is infinite and A80: B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } and A81: f . (n + 1) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } ),fn) by A77; set BBX = { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; set BJ = { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } ; set BC = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } ; consider BD being set such that A82: BD is_Dickson-basis_of N,R and A83: BD is finite by A1, Def10; BD in Dickson-bases (N,R) by A1, A82, Def13; then BD in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A83; then B in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A80; then A84: ex BB being Element of Dickson-bases (N,R) st ( B = BB & BB is finite ) ; then A85: B is_Dickson-basis_of N,R by A1, Def13; then A86: B c= N by Def9; now__::_thesis:_ex_b_being_Element_of_rng_s_st_ (_b_in_B_&_not_H1(b,fn)_is_finite_) A87: B is finite by A84; assume A88: for b being Element of rng s st b in B holds H1(b,fn) is finite ; ::_thesis: contradiction deffunc H4( Element of rng s) -> set = H1($1,fn); set Ball = { H4(b) where b is Element of rng s : b in B } ; set iN = { w where w is Element of NAT : ( s . fn <= s . w & fn < w ) } ; A89: { H4(b) where b is Element of rng s : b in B } is finite from FRAENKEL:sch_21(A87); now__::_thesis:_for_X_being_set_st_X_in__{__H4(b)_where_b_is_Element_of_rng_s_:_b_in_B__}__holds_ X_is_finite let X be set ; ::_thesis: ( X in { H4(b) where b is Element of rng s : b in B } implies X is finite ) assume X in { H4(b) where b is Element of rng s : b in B } ; ::_thesis: X is finite then ex b being Element of rng s st ( X = H1(b,fn) & b in B ) ; hence X is finite by A88; ::_thesis: verum end; then A90: union { H4(b) where b is Element of rng s : b in B } is finite by A89, FINSET_1:7; { w where w is Element of NAT : ( s . fn <= s . w & fn < w ) } c= union { H4(b) where b is Element of rng s : b in B } proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { w where w is Element of NAT : ( s . fn <= s . w & fn < w ) } or x in union { H4(b) where b is Element of rng s : b in B } ) assume x in { w where w is Element of NAT : ( s . fn <= s . w & fn < w ) } ; ::_thesis: x in union { H4(b) where b is Element of rng s : b in B } then consider w being Element of NAT such that A91: x = w and A92: s . fn <= s . w and A93: f . n < w ; A94: s . w in N by A78, A92, A93; reconsider sw = s . w as Element of R ; consider b being Element of R such that A95: b in B and A96: b <= sw by A85, A94, Def9; A97: B c= N by A85, Def9; N c= rng s proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in N or x in rng s ) assume x in N ; ::_thesis: x in rng s then A98: ex u being Element of NAT st ( x = s . u & s . fn <= s . u & fn < u ) by A78; dom s = NAT by FUNCT_2:def_1; hence x in rng s by A98, FUNCT_1:3; ::_thesis: verum end; then B c= rng s by A97, XBOOLE_1:1; then reconsider b = b as Element of rng s by A95; A99: w in H1(b,fn) by A93, A96; H1(b,fn) in { H4(b) where b is Element of rng s : b in B } by A95; hence x in union { H4(b) where b is Element of rng s : b in B } by A91, A99, TARSKI:def_4; ::_thesis: verum end; hence contradiction by A79, A90; ::_thesis: verum end; then consider b being Element of rng s such that A100: b in B and A101: H1(b,fn) is infinite ; b in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } by A100, A101; then choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } ; then ex b being Element of B st ( choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } = b & ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) ) ; then choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } in N by A86, TARSKI:def_3; then ex j being Element of NAT st ( choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,fn) is infinite ) } = s . j & s . fn <= s . j & fn < j ) by A78; hence f . n < f . (n + 1) by A81, Def12; ::_thesis: verum end; hence f is increasing by SEQM_3:def_6; ::_thesis: for n being Element of NAT holds f . n is Element of NAT let n be Element of NAT ; ::_thesis: f . n is Element of NAT f . n in rng f by A16, FUNCT_1:def_3; hence f . n is Element of NAT by A17; ::_thesis: verum end; then reconsider f = f as increasing sequence of NAT ; s * f = s * f ; hence s * f is subsequence of s ; ::_thesis: s * f is weakly-ascending let n be Element of NAT ; :: according to DICKSON:def_2 ::_thesis: [((s * f) . n),((s * f) . (n + 1))] in the InternalRel of R A102: (s * f) . n = s . (f . n) by A16, FUNCT_1:13; A103: (s * f) . (n + 1) = s . (f . (n + 1)) by A16, FUNCT_1:13; consider N being Subset of the carrier of R, B being non empty Subset of the carrier of R such that A104: N = { (s . w) where w is Element of NAT : ( s . (f . n) <= s . w & f . n < w ) } and A105: { w where w is Element of NAT : ( s . (f . n) <= s . w & f . n < w ) } is infinite and A106: B = choose { BB where BB is Element of Dickson-bases (N,R) : BB is finite } and A107: f . (n + 1) = s mindex ((choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } ),(f . n)) by A77; set BX = { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } ; set sfn1 = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } ; set BBX = { BB where BB is Element of Dickson-bases (N,R) : BB is finite } ; consider BD being set such that A108: BD is_Dickson-basis_of N,R and A109: BD is finite by A1, Def10; BD in Dickson-bases (N,R) by A1, A108, Def13; then BD in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A109; then B in { BB where BB is Element of Dickson-bases (N,R) : BB is finite } by A106; then A110: ex BB being Element of Dickson-bases (N,R) st ( BB = B & BB is finite ) ; then A111: B is_Dickson-basis_of N,R by A1, Def13; now__::_thesis:_ex_b_being_Element_of_rng_s_st_ (_b_in_B_&_not_H1(b,f_._n)_is_finite_) A112: B is finite by A110; assume A113: for b being Element of rng s st b in B holds H1(b,f . n) is finite ; ::_thesis: contradiction deffunc H4( Element of rng s) -> set = H1($1,f . n); set Ball = { H4(b) where b is Element of rng s : b in B } ; set iN = { w where w is Element of NAT : ( s . (f . n) <= s . w & f . n < w ) } ; A114: { H4(b) where b is Element of rng s : b in B } is finite from FRAENKEL:sch_21(A112); now__::_thesis:_for_X_being_set_st_X_in__{__H4(b)_where_b_is_Element_of_rng_s_:_b_in_B__}__holds_ X_is_finite let X be set ; ::_thesis: ( X in { H4(b) where b is Element of rng s : b in B } implies X is finite ) assume X in { H4(b) where b is Element of rng s : b in B } ; ::_thesis: X is finite then ex b being Element of rng s st ( X = H1(b,f . n) & b in B ) ; hence X is finite by A113; ::_thesis: verum end; then A115: union { H4(b) where b is Element of rng s : b in B } is finite by A114, FINSET_1:7; { w where w is Element of NAT : ( s . (f . n) <= s . w & f . n < w ) } c= union { H4(b) where b is Element of rng s : b in B } proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { w where w is Element of NAT : ( s . (f . n) <= s . w & f . n < w ) } or x in union { H4(b) where b is Element of rng s : b in B } ) assume x in { w where w is Element of NAT : ( s . (f . n) <= s . w & f . n < w ) } ; ::_thesis: x in union { H4(b) where b is Element of rng s : b in B } then consider w being Element of NAT such that A116: x = w and A117: s . (f . n) <= s . w and A118: f . n < w ; A119: s . w in N by A104, A117, A118; reconsider sw = s . w as Element of R ; consider b being Element of R such that A120: b in B and A121: b <= sw by A111, A119, Def9; A122: B c= N by A111, Def9; N c= rng s proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in N or x in rng s ) assume x in N ; ::_thesis: x in rng s then A123: ex u being Element of NAT st ( x = s . u & s . (f . n) <= s . u & f . n < u ) by A104; dom s = NAT by FUNCT_2:def_1; hence x in rng s by A123, FUNCT_1:3; ::_thesis: verum end; then B c= rng s by A122, XBOOLE_1:1; then reconsider b = b as Element of rng s by A120; A124: w in H1(b,f . n) by A118, A121; H1(b,f . n) in { H4(b) where b is Element of rng s : b in B } by A120; hence x in union { H4(b) where b is Element of rng s : b in B } by A116, A124, TARSKI:def_4; ::_thesis: verum end; hence contradiction by A105, A115; ::_thesis: verum end; then consider b being Element of rng s such that A125: b in B and A126: H1(b,f . n) is infinite ; b in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } by A125, A126; then choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } in { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } ; then ex b being Element of B st ( b = choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } & ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) ) ; then A127: choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } in B ; B c= N by A111, Def9; then choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } in N by A127; then ex w being Element of NAT st ( choose { b where b is Element of B : ex b9 being Element of rng s st ( b9 = b & H1(b9,f . n) is infinite ) } = s . w & s . (f . n) <= s . w & f . n < w ) by A104; then (s * f) . n <= (s * f) . (n + 1) by A102, A103, A107, Def12; hence [((s * f) . n),((s * f) . (n + 1))] in the InternalRel of R by ORDERS_2:def_5; ::_thesis: verum end; theorem Th36: :: DICKSON:36 for R being RelStr st R is empty holds R is Dickson proof let R be RelStr ; ::_thesis: ( R is empty implies R is Dickson ) assume A1: R is empty ; ::_thesis: R is Dickson now__::_thesis:_for_N_being_Subset_of_R_ex_B_being_set_st_ (_B_is_Dickson-basis_of_N,R_&_B_is_finite_) let N be Subset of R; ::_thesis: ex B being set st ( B is_Dickson-basis_of N,R & B is finite ) take B = {} ; ::_thesis: ( B is_Dickson-basis_of N,R & B is finite ) N = {} the carrier of R by A1; hence B is_Dickson-basis_of N,R by Th25; ::_thesis: B is finite thus B is finite ; ::_thesis: verum end; hence R is Dickson by Def10; ::_thesis: verum end; theorem Th37: :: DICKSON:37 for M, N being RelStr st M is Dickson & N is Dickson & M is quasi_ordered & N is quasi_ordered holds ( [:M,N:] is quasi_ordered & [:M,N:] is Dickson ) proof let M, N be RelStr ; ::_thesis: ( M is Dickson & N is Dickson & M is quasi_ordered & N is quasi_ordered implies ( [:M,N:] is quasi_ordered & [:M,N:] is Dickson ) ) assume that A1: M is Dickson and A2: N is Dickson and A3: M is quasi_ordered and A4: N is quasi_ordered ; ::_thesis: ( [:M,N:] is quasi_ordered & [:M,N:] is Dickson ) reconsider M9 = M as reflexive transitive RelStr by A3, Def3; reconsider N9 = N as reflexive transitive RelStr by A4, Def3; [:M9,N9:] is reflexive ; hence A5: [:M,N:] is quasi_ordered by Def3; ::_thesis: [:M,N:] is Dickson percases ( ( not M is empty & not N is empty ) or M is empty or N is empty ) ; suppose ( not M is empty & not N is empty ) ; ::_thesis: [:M,N:] is Dickson then reconsider Me = M, Ne = N as non empty RelStr ; set CPMN = [:Me,Ne:]; for f being sequence of [:Me,Ne:] ex i, j being Element of NAT st ( i < j & f . i <= f . j ) proof let f be sequence of [:Me,Ne:]; ::_thesis: ex i, j being Element of NAT st ( i < j & f . i <= f . j ) deffunc H1( Element of NAT ) -> Element of the carrier of Me = (f . $1) `1 ; consider a being Function of NAT, the carrier of Me such that A6: for x being Element of NAT holds a . x = H1(x) from FUNCT_2:sch_4(); reconsider a = a as sequence of Me ; consider sa being sequence of Me such that A7: sa is subsequence of a and A8: sa is weakly-ascending by A1, Th35; consider NS being increasing sequence of NAT such that A9: sa = a * NS by A7, VALUED_0:def_17; deffunc H2( Element of NAT ) -> Element of the carrier of Ne = (f . (NS . $1)) `2 ; consider b being Function of NAT, the carrier of Ne such that A10: for x being Element of NAT holds b . x = H2(x) from FUNCT_2:sch_4(); reconsider b = b as sequence of Ne ; consider i, j being Element of NAT such that A11: i < j and A12: b . i <= b . j by A2, Th29; take NS . i ; ::_thesis: ex j being Element of NAT st ( NS . i < j & f . (NS . i) <= f . j ) take NS . j ; ::_thesis: ( NS . i < NS . j & f . (NS . i) <= f . (NS . j) ) dom NS = NAT by FUNCT_2:def_1; hence NS . i < NS . j by A11, VALUED_0:def_13; ::_thesis: f . (NS . i) <= f . (NS . j) reconsider x = f . (NS . i), y = f . (NS . j) as Element of [:Me,Ne:] ; A13: dom sa = NAT by FUNCT_2:def_1; then A14: sa . i = a . (NS . i) by A9, FUNCT_1:12 .= (f . (NS . i)) `1 by A6 ; A15: sa . j = a . (NS . j) by A9, A13, FUNCT_1:12 .= (f . (NS . j)) `1 by A6 ; M is transitive by A3, Def3; then A16: x `1 <= y `1 by A8, A11, A14, A15, Th4; A17: b . i = x `2 by A10; b . j = y `2 by A10; hence f . (NS . i) <= f . (NS . j) by A12, A16, A17, YELLOW_3:12; ::_thesis: verum end; then for N being non empty Subset of [:Me,Ne:] holds ( min-classes N is finite & not min-classes N is empty ) by A5, Th31; hence [:M,N:] is Dickson by A5, Th32; ::_thesis: verum end; supposeA18: ( M is empty or N is empty ) ; ::_thesis: [:M,N:] is Dickson now__::_thesis:_[:_the_carrier_of_M,_the_carrier_of_N:]_is_empty percases ( M is empty or N is empty ) by A18; suppose M is empty ; ::_thesis: [: the carrier of M, the carrier of N:] is empty then reconsider M2 = the carrier of M as empty set ; [:M2, the carrier of N:] is empty ; hence [: the carrier of M, the carrier of N:] is empty ; ::_thesis: verum end; suppose N is empty ; ::_thesis: [: the carrier of M, the carrier of N:] is empty then reconsider N2 = the carrier of N as empty set ; [: the carrier of M,N2:] is empty ; hence [: the carrier of M, the carrier of N:] is empty ; ::_thesis: verum end; end; end; then [:M,N:] is empty by YELLOW_3:def_2; hence [:M,N:] is Dickson by Th36; ::_thesis: verum end; end; end; theorem Th38: :: DICKSON:38 for R, S being RelStr st R,S are_isomorphic & R is Dickson & R is quasi_ordered holds ( S is quasi_ordered & S is Dickson ) proof let R, S be RelStr ; ::_thesis: ( R,S are_isomorphic & R is Dickson & R is quasi_ordered implies ( S is quasi_ordered & S is Dickson ) ) assume that A1: R,S are_isomorphic and A2: R is Dickson and A3: R is quasi_ordered ; ::_thesis: ( S is quasi_ordered & S is Dickson ) set CRS = the carrier of S; set IRS = the InternalRel of S; percases ( ( not R is empty & not S is empty ) or R is empty or S is empty ) ; suppose ( not R is empty & not S is empty ) ; ::_thesis: ( S is quasi_ordered & S is Dickson ) then reconsider Re = R, Se = S as non empty RelStr ; consider f being Function of Re,Se such that A4: f is isomorphic by A1, WAYBEL_1:def_8; A5: f is V7() by A4, WAYBEL_0:66; A6: rng f = the carrier of Se by A4, WAYBEL_0:66; A7: Re is reflexive by A3, Def3; A8: Re is transitive by A3, Def3; A9: Se is reflexive by A1, A7, WAYBEL20:15; Se is transitive by A1, A8, WAYBEL20:16; hence A10: S is quasi_ordered by A9, Def3; ::_thesis: S is Dickson now__::_thesis:_for_t_being_sequence_of_Se_ex_i,_j_being_Element_of_NAT_st_ (_i_<_j_&_t_._i_<=_t_._j_) let t be sequence of Se; ::_thesis: ex i, j being Element of NAT st ( i < j & t . i <= t . j ) reconsider fi = f " as Function of the carrier of Se, the carrier of Re by A5, A6, FUNCT_2:25; deffunc H1( Element of NAT ) -> Element of the carrier of Re = fi . (t . $1); consider sr being Function of NAT, the carrier of Re such that A11: for x being Element of NAT holds sr . x = H1(x) from FUNCT_2:sch_4(); reconsider sr = sr as sequence of Re ; consider i, j being Element of NAT such that A12: i < j and A13: sr . i <= sr . j by A2, Th29; take i = i; ::_thesis: ex j being Element of NAT st ( i < j & t . i <= t . j ) take j = j; ::_thesis: ( i < j & t . i <= t . j ) thus i < j by A12; ::_thesis: t . i <= t . j A14: f . (sr . i) = f . ((f ") . (t . i)) by A11 .= t . i by A5, A6, FUNCT_1:35 ; f . (sr . j) = f . ((f ") . (t . j)) by A11 .= t . j by A5, A6, FUNCT_1:35 ; hence t . i <= t . j by A4, A13, A14, WAYBEL_0:66; ::_thesis: verum end; then for N being non empty Subset of Se holds ( min-classes N is finite & not min-classes N is empty ) by A10, Th31; hence S is Dickson by A10, Th32; ::_thesis: verum end; supposeA15: ( R is empty or S is empty ) ; ::_thesis: ( S is quasi_ordered & S is Dickson ) A16: now__::_thesis:_S_is_empty percases ( R is empty or S is empty ) by A15; supposeA17: R is empty ; ::_thesis: S is empty ex f being Function of R,S st f is isomorphic by A1, WAYBEL_1:def_8; hence S is empty by A17, WAYBEL_0:def_38; ::_thesis: verum end; suppose S is empty ; ::_thesis: S is empty hence S is empty ; ::_thesis: verum end; end; end; then for x being set st x in the carrier of S holds [x,x] in the InternalRel of S ; then A18: the InternalRel of S is_reflexive_in the carrier of S by RELAT_2:def_1; for x, y, z being set st x in the carrier of S & y in the carrier of S & z in the carrier of S & [x,y] in the InternalRel of S & [y,z] in the InternalRel of S holds [x,z] in the InternalRel of S by A16; then A19: the InternalRel of S is_transitive_in the carrier of S by RELAT_2:def_8; A20: S is reflexive by A18, ORDERS_2:def_2; S is transitive by A19, ORDERS_2:def_3; hence S is quasi_ordered by A20, Def3; ::_thesis: S is Dickson thus S is Dickson by A16, Th36; ::_thesis: verum end; end; end; theorem Th39: :: DICKSON:39 for p being RelStr-yielding ManySortedSet of 1 for z being Element of 1 holds p . z, product p are_isomorphic proof let p be RelStr-yielding ManySortedSet of 1; ::_thesis: for z being Element of 1 holds p . z, product p are_isomorphic let z be Element of 1; ::_thesis: p . z, product p are_isomorphic deffunc H1( set ) -> set = 0 .--> $1; consider f being Function such that A1: dom f = the carrier of (p . z) and A2: for x being set st x in the carrier of (p . z) holds f . x = H1(x) from FUNCT_1:sch_3(); A3: z = 0 by CARD_1:49, TARSKI:def_1; A4: 0 in 1 by CARD_1:49, TARSKI:def_1; now__::_thesis:_for_x_being_set_st_x_in_the_carrier_of_(p_._z)_holds_ f_._x_in_the_carrier_of_(product_p) let x be set ; ::_thesis: ( x in the carrier of (p . z) implies f . x in the carrier of (product p) ) assume A5: x in the carrier of (p . z) ; ::_thesis: f . x in the carrier of (product p) then A6: f . x = 0 .--> x by A2; set g = 0 .--> x; A7: dom (0 .--> x) = {0} by FUNCOP_1:13 .= dom (Carrier p) by CARD_1:49, PARTFUN1:def_2 ; now__::_thesis:_for_y_being_set_st_y_in_dom_(Carrier_p)_holds_ (0_.-->_x)_._y_in_(Carrier_p)_._y let y be set ; ::_thesis: ( y in dom (Carrier p) implies (0 .--> x) . y in (Carrier p) . y ) assume y in dom (Carrier p) ; ::_thesis: (0 .--> x) . y in (Carrier p) . y then A8: y in 1 ; then A9: y = 0 by CARD_1:49, TARSKI:def_1; ex R being 1-sorted st ( R = p . y & (Carrier p) . y = the carrier of R ) by A8, PRALG_1:def_13; hence (0 .--> x) . y in (Carrier p) . y by A3, A5, A9, FUNCOP_1:72; ::_thesis: verum end; then f . x in product (Carrier p) by A6, A7, CARD_3:def_5; hence f . x in the carrier of (product p) by YELLOW_1:def_4; ::_thesis: verum end; then reconsider f = f as Function of (p . z),(product p) by A1, FUNCT_2:3; now__::_thesis:_f_is_isomorphic percases ( p . z is empty or not p . z is empty ) ; supposeA10: p . z is empty ; ::_thesis: f is isomorphic now__::_thesis:_the_carrier_of_(product_p)_is_empty assume not the carrier of (product p) is empty ; ::_thesis: contradiction then A11: not product (Carrier p) is empty by YELLOW_1:def_4; set x = the Element of product (Carrier p); A12: ex g being Function st ( the Element of product (Carrier p) = g & dom g = dom (Carrier p) & ( for y being set st y in dom (Carrier p) holds g . y in (Carrier p) . y ) ) by A11, CARD_3:def_5; A13: 0 in dom (Carrier p) by A4, PARTFUN1:def_2; consider R being 1-sorted such that A14: R = p . 0 and A15: (Carrier p) . 0 = the carrier of R by A4, PRALG_1:def_13; R is empty by A10, A14, CARD_1:49, TARSKI:def_1; hence contradiction by A12, A13, A15; ::_thesis: verum end; then product p is empty ; hence f is isomorphic by A10, WAYBEL_0:def_38; ::_thesis: verum end; supposeA16: not p . z is empty ; ::_thesis: f is isomorphic then reconsider pz = p . z as non empty RelStr ; now__::_thesis:_for_i_being_set_st_i_in_rng_p_holds_ i_is_non_empty_1-sorted let i be set ; ::_thesis: ( i in rng p implies i is non empty 1-sorted ) assume i in rng p ; ::_thesis: i is non empty 1-sorted then consider x being set such that A17: x in dom p and A18: i = p . x by FUNCT_1:def_3; x in 1 by A17; then i = p . 0 by A18, CARD_1:49, TARSKI:def_1; hence i is non empty 1-sorted by A16, CARD_1:49, TARSKI:def_1; ::_thesis: verum end; then reconsider np = p as RelStr-yielding yielding_non-empty_carriers ManySortedSet of 1 by YELLOW_6:def_2; not product np is empty ; then reconsider pp = product p as non empty RelStr ; now__::_thesis:_for_x1,_x2_being_set_st_x1_in_dom_f_&_x2_in_dom_f_&_f_._x1_=_f_._x2_holds_ x1_=_x2 let x1, x2 be set ; ::_thesis: ( x1 in dom f & x2 in dom f & f . x1 = f . x2 implies x1 = x2 ) assume that A19: x1 in dom f and A20: x2 in dom f and A21: f . x1 = f . x2 ; ::_thesis: x1 = x2 A22: f . x1 = 0 .--> x1 by A2, A19 .= [:{0},{x1}:] ; A23: f . x2 = 0 .--> x2 by A2, A20 .= [:{0},{x2}:] ; A24: 0 in {0} by TARSKI:def_1; x1 in {x1} by TARSKI:def_1; then [0,x1] in f . x2 by A21, A22, A24, ZFMISC_1:def_2; then ex xa, ya being set st ( xa in {0} & ya in {x2} & [0,x1] = [xa,ya] ) by A23, ZFMISC_1:def_2; then x1 in {x2} by XTUPLE_0:1; hence x1 = x2 by TARSKI:def_1; ::_thesis: verum end; then A25: f is one-to-one by FUNCT_1:def_4; now__::_thesis:_for_y_being_set_holds_ (_(_y_in_rng_f_implies_y_in_the_carrier_of_(product_p)_)_&_(_y_in_the_carrier_of_(product_p)_implies_y_in_rng_f_)_) let y be set ; ::_thesis: ( ( y in rng f implies y in the carrier of (product p) ) & ( y in the carrier of (product p) implies y in rng f ) ) thus ( y in rng f implies y in the carrier of (product p) ) ; ::_thesis: ( y in the carrier of (product p) implies y in rng f ) assume y in the carrier of (product p) ; ::_thesis: y in rng f then y in product (Carrier p) by YELLOW_1:def_4; then consider g being Function such that A26: y = g and A27: dom g = dom (Carrier p) and A28: for x being set st x in dom (Carrier p) holds g . x in (Carrier p) . x by CARD_3:def_5; A29: dom g = {0} by A27, CARD_1:49, PARTFUN1:def_2; A30: 0 in dom (Carrier p) by A4, PARTFUN1:def_2; A31: ex R being 1-sorted st ( R = p . 0 & (Carrier p) . 0 = the carrier of R ) by A4, PRALG_1:def_13; A32: g = 0 .--> (g . 0) by A29, Th1; A33: f . (g . 0) = 0 .--> (g . 0) by A2, A3, A28, A30, A31; g . 0 in dom f by A1, A3, A28, A30, A31; hence y in rng f by A26, A32, A33, FUNCT_1:def_3; ::_thesis: verum end; then A34: rng f = the carrier of (product p) by TARSKI:1; reconsider f9 = f as Function of pz,pp ; now__::_thesis:_for_x,_y_being_Element_of_pz_holds_ (_(_x_<=_y_implies_f9_._x_<=_f9_._y_)_&_(_f9_._x_<=_f9_._y_implies_x_<=_y_)_) let x, y be Element of pz; ::_thesis: ( ( x <= y implies f9 . x <= f9 . y ) & ( f9 . x <= f9 . y implies x <= y ) ) not product np is empty ; then A35: not product (Carrier p) is empty by YELLOW_1:def_4; A36: f9 . x is Element of product (Carrier p) by YELLOW_1:def_4; hereby ::_thesis: ( f9 . x <= f9 . y implies x <= y ) assume A37: x <= y ; ::_thesis: f9 . x <= f9 . y ex f1, g1 being Function st ( f1 = f9 . x & g1 = f9 . y & ( for i being set st i in 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) ) ) proof set f1 = 0 .--> x; set g1 = 0 .--> y; reconsider f1 = 0 .--> x as Function ; reconsider g1 = 0 .--> y as Function ; take f1 ; ::_thesis: ex g1 being Function st ( f1 = f9 . x & g1 = f9 . y & ( for i being set st i in 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) ) ) take g1 ; ::_thesis: ( f1 = f9 . x & g1 = f9 . y & ( for i being set st i in 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) ) ) thus f1 = f9 . x by A2; ::_thesis: ( g1 = f9 . y & ( for i being set st i in 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) ) ) thus g1 = f9 . y by A2; ::_thesis: for i being set st i in 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) let i be set ; ::_thesis: ( i in 1 implies ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) ) assume A38: i in 1 ; ::_thesis: ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) A39: i = 0 by A38, CARD_1:49, TARSKI:def_1; 0 in dom p by A4, PARTFUN1:def_2; then p . 0 in rng p by FUNCT_1:def_3; then reconsider p0 = p . 0 as RelStr by YELLOW_1:def_3; set R = p0; reconsider x9 = x as Element of p0 by CARD_1:49, TARSKI:def_1; reconsider y9 = y as Element of p0 by CARD_1:49, TARSKI:def_1; take p0 ; ::_thesis: ex xi, yi being Element of p0 st ( p0 = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) take x9 ; ::_thesis: ex yi being Element of p0 st ( p0 = p . i & x9 = f1 . i & yi = g1 . i & x9 <= yi ) take y9 ; ::_thesis: ( p0 = p . i & x9 = f1 . i & y9 = g1 . i & x9 <= y9 ) thus p0 = p . i by A38, CARD_1:49, TARSKI:def_1; ::_thesis: ( x9 = f1 . i & y9 = g1 . i & x9 <= y9 ) thus x9 = f1 . i by A39, FUNCOP_1:72; ::_thesis: ( y9 = g1 . i & x9 <= y9 ) thus y9 = g1 . i by A39, FUNCOP_1:72; ::_thesis: x9 <= y9 thus x9 <= y9 by A37, CARD_1:49, TARSKI:def_1; ::_thesis: verum end; hence f9 . x <= f9 . y by A35, A36, YELLOW_1:def_4; ::_thesis: verum end; assume f9 . x <= f9 . y ; ::_thesis: x <= y then consider f1, g1 being Function such that A40: f1 = f9 . x and A41: g1 = f9 . y and A42: for i being set st i in 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) by A35, A36, YELLOW_1:def_4; consider R being RelStr , xi, yi being Element of R such that A43: R = p . 0 and A44: xi = f1 . 0 and A45: yi = g1 . 0 and A46: xi <= yi by A4, A42; f1 = 0 .--> x by A2, A40; then A47: xi = x by A44, FUNCOP_1:72; A48: g1 = 0 .--> y by A2, A41; R = pz by A43, CARD_1:49, TARSKI:def_1; hence x <= y by A45, A46, A47, A48, FUNCOP_1:72; ::_thesis: verum end; hence f is isomorphic by A25, A34, WAYBEL_0:66; ::_thesis: verum end; end; end; hence p . z, product p are_isomorphic by WAYBEL_1:def_8; ::_thesis: verum end; registration let X be set ; let p be RelStr-yielding ManySortedSet of X; let Y be Subset of X; clusterp | Y -> RelStr-yielding ; coherence p | Y is RelStr-yielding proof now__::_thesis:_for_v_being_set_st_v_in_rng_(p_|_Y)_holds_ v_is_RelStr let v be set ; ::_thesis: ( v in rng (p | Y) implies v is RelStr ) assume v in rng (p | Y) ; ::_thesis: v is RelStr then consider x being set such that A1: x in dom (p | Y) and A2: v = (p | Y) . x by FUNCT_1:def_3; A3: x in Y by A1; then A4: (p | Y) . x = p . x by FUNCT_1:49; x in X by A3; then x in dom p by PARTFUN1:def_2; then p . x in rng p by FUNCT_1:3; hence v is RelStr by A2, A4, YELLOW_1:def_3; ::_thesis: verum end; hence p | Y is RelStr-yielding by YELLOW_1:def_3; ::_thesis: verum end; end; theorem Th40: :: DICKSON:40 for n being non empty Nat for p being RelStr-yielding ManySortedSet of n holds ( not product p is empty iff p is non-Empty ) proof let n be non empty Nat; ::_thesis: for p being RelStr-yielding ManySortedSet of n holds ( not product p is empty iff p is non-Empty ) let p be RelStr-yielding ManySortedSet of n; ::_thesis: ( not product p is empty iff p is non-Empty ) hereby ::_thesis: ( p is non-Empty implies not product p is empty ) assume not product p is empty ; ::_thesis: p is non-Empty then not product (Carrier p) is empty by YELLOW_1:def_4; then consider z being set such that A1: z in product (Carrier p) by XBOOLE_0:def_1; A2: ex g being Function st ( z = g & dom g = dom (Carrier p) & ( for i being set st i in dom (Carrier p) holds g . i in (Carrier p) . i ) ) by A1, CARD_3:def_5; now__::_thesis:_for_S_being_1-sorted_st_S_in_rng_p_holds_ not_S_is_empty let S be 1-sorted ; ::_thesis: ( S in rng p implies not S is empty ) assume S in rng p ; ::_thesis: not S is empty then consider x being set such that A3: x in dom p and A4: S = p . x by FUNCT_1:def_3; A5: x in n by A3; then A6: x in dom (Carrier p) by PARTFUN1:def_2; ex R being 1-sorted st ( R = p . x & (Carrier p) . x = the carrier of R ) by A5, PRALG_1:def_13; hence not S is empty by A2, A4, A6; ::_thesis: verum end; hence p is non-Empty by WAYBEL_3:def_7; ::_thesis: verum end; assume A7: p is non-Empty ; ::_thesis: not product p is empty A8: dom (Carrier p) = n by PARTFUN1:def_2; deffunc H1( set ) -> Element of (Carrier p) . $1 = choose ((Carrier p) . $1); consider g being Function such that A9: dom g = dom (Carrier p) and A10: for i being set st i in dom (Carrier p) holds g . i = H1(i) from FUNCT_1:sch_3(); set x = g; now__::_thesis:_ex_g_being_Function_st_ (_g_=_g_&_dom_g_=_dom_(Carrier_p)_&_(_for_i_being_set_st_i_in_dom_(Carrier_p)_holds_ g_._i_in_(Carrier_p)_._i_)_) take g = g; ::_thesis: ( g = g & dom g = dom (Carrier p) & ( for i being set st i in dom (Carrier p) holds g . i in (Carrier p) . i ) ) thus g = g ; ::_thesis: ( dom g = dom (Carrier p) & ( for i being set st i in dom (Carrier p) holds g . i in (Carrier p) . i ) ) thus dom g = dom (Carrier p) by A9; ::_thesis: for i being set st i in dom (Carrier p) holds g . i in (Carrier p) . i thus for i being set st i in dom (Carrier p) holds g . i in (Carrier p) . i ::_thesis: verum proof let i be set ; ::_thesis: ( i in dom (Carrier p) implies g . i in (Carrier p) . i ) assume A11: i in dom (Carrier p) ; ::_thesis: g . i in (Carrier p) . i i in dom p by A8, A11, PARTFUN1:def_2; then A12: p . i in rng p by FUNCT_1:def_3; then reconsider pi1 = p . i as RelStr by YELLOW_1:def_3; not pi1 is empty by A7, A12, WAYBEL_3:def_7; then A13: not the carrier of pi1 is empty ; i in n by A11; then A14: ex R being 1-sorted st ( R = p . i & (Carrier p) . i = the carrier of R ) by PRALG_1:def_13; g . i = choose ((Carrier p) . i) by A10, A11; hence g . i in (Carrier p) . i by A13, A14; ::_thesis: verum end; end; then not product (Carrier p) is empty by CARD_3:def_5; hence not product p is empty by YELLOW_1:def_4; ::_thesis: verum end; theorem Th41: :: DICKSON:41 for n being non empty Nat for p being RelStr-yielding ManySortedSet of n + 1 for ns being Subset of (n + 1) for ne being Element of n + 1 st ns = n & ne = n holds [:(product (p | ns)),(p . ne):], product p are_isomorphic proof let n be non empty Nat; ::_thesis: for p being RelStr-yielding ManySortedSet of n + 1 for ns being Subset of (n + 1) for ne being Element of n + 1 st ns = n & ne = n holds [:(product (p | ns)),(p . ne):], product p are_isomorphic let p be RelStr-yielding ManySortedSet of n + 1; ::_thesis: for ns being Subset of (n + 1) for ne being Element of n + 1 st ns = n & ne = n holds [:(product (p | ns)),(p . ne):], product p are_isomorphic let ns be Subset of (n + 1); ::_thesis: for ne being Element of n + 1 st ns = n & ne = n holds [:(product (p | ns)),(p . ne):], product p are_isomorphic let ne be Element of n + 1; ::_thesis: ( ns = n & ne = n implies [:(product (p | ns)),(p . ne):], product p are_isomorphic ) assume that A1: ns = n and A2: ne = n ; ::_thesis: [:(product (p | ns)),(p . ne):], product p are_isomorphic set S = [:(product (p | ns)),(p . ne):]; set T = product p; set CP = [: the carrier of (product (p | ns)), the carrier of (p . ne):]; set CPP = the carrier of (product p); A3: dom (Carrier (p | ns)) = n by A1, PARTFUN1:def_2; percases ( the carrier of (product p) is empty or not the carrier of (product p) is empty ) ; supposeA4: the carrier of (product p) is empty ; ::_thesis: [:(product (p | ns)),(p . ne):], product p are_isomorphic then A5: product p is empty ; not p is non-Empty by A4; then consider r1 being 1-sorted such that A6: r1 in rng p and A7: r1 is empty by WAYBEL_3:def_7; consider x being set such that A8: x in dom p and A9: r1 = p . x by A6, FUNCT_1:def_3; x in n + 1 by A8; then A10: x in n \/ {n} by AFINSQ_1:2; A11: [:(product (p | ns)),(p . ne):] is empty proof percases ( x in n or x in {n} ) by A10, XBOOLE_0:def_3; supposeA12: x in n ; ::_thesis: [:(product (p | ns)),(p . ne):] is empty then A13: (p | ns) . x = p . x by A1, FUNCT_1:49; x in dom (p | ns) by A1, A12, PARTFUN1:def_2; then p . x in rng (p | ns) by A13, FUNCT_1:def_3; then not p | ns is non-Empty by A7, A9, WAYBEL_3:def_7; then reconsider ptemp = the carrier of (product (p | ns)) as empty set by A1, Th40; [:ptemp, the carrier of (p . ne):] is empty ; hence [:(product (p | ns)),(p . ne):] is empty by YELLOW_3:def_2; ::_thesis: verum end; suppose x in {n} ; ::_thesis: [:(product (p | ns)),(p . ne):] is empty then p . ne is empty by A2, A7, A9, TARSKI:def_1; then reconsider pne = the carrier of (p . ne) as empty set ; reconsider ptemp = the carrier of (product (p | ns)) as set ; [:ptemp,pne:] is empty ; hence [:(product (p | ns)),(p . ne):] is empty by YELLOW_3:def_2; ::_thesis: verum end; end; end; then A14: dom {} = the carrier of [:(product (p | ns)),(p . ne):] ; for x being set st x in {} holds {} . x in the carrier of (product p) ; then reconsider f = {} as Function of [:(product (p | ns)),(p . ne):],(product p) by A14, FUNCT_2:3; f is isomorphic by A5, A11, WAYBEL_0:def_38; hence [:(product (p | ns)),(p . ne):], product p are_isomorphic by WAYBEL_1:def_8; ::_thesis: verum end; supposeA15: not the carrier of (product p) is empty ; ::_thesis: [:(product (p | ns)),(p . ne):], product p are_isomorphic then reconsider CPP = the carrier of (product p) as non empty set ; reconsider T = product p as non empty RelStr by A15; A16: p is non-Empty by A15, Th40; reconsider p9 = p as RelStr-yielding non-Empty ManySortedSet of n + 1 by A15, Th40; not p9 . ne is empty ; then reconsider pne2 = the carrier of (p . ne) as non empty set ; now__::_thesis:_for_S_being_1-sorted_st_S_in_rng_(p_|_ns)_holds_ not_S_is_empty let S be 1-sorted ; ::_thesis: ( S in rng (p | ns) implies not S is empty ) assume S in rng (p | ns) ; ::_thesis: not S is empty then consider x being set such that A17: x in dom (p | ns) and A18: S = (p | ns) . x by FUNCT_1:def_3; x in ns by A17; then x in n + 1 ; then A19: x in dom p by PARTFUN1:def_2; S = p . x by A17, A18, FUNCT_1:47; then S in rng p by A19, FUNCT_1:def_3; hence not S is empty by A16, WAYBEL_3:def_7; ::_thesis: verum end; then A20: p | ns is non-Empty by WAYBEL_3:def_7; then A21: not product (p | ns) is empty ; reconsider ppns2 = the carrier of (product (p | ns)) as non empty set by A20; [: the carrier of (product (p | ns)), the carrier of (p . ne):] = [:ppns2,pne2:] ; then reconsider S = [:(product (p | ns)),(p . ne):] as non empty RelStr by YELLOW_3:def_2; [: the carrier of (product (p | ns)), the carrier of (p . ne):] = the carrier of S by YELLOW_3:def_2; then reconsider CP9 = [: the carrier of (product (p | ns)), the carrier of (p . ne):] as non empty set ; defpred S1[ set , set ] means ex a being Function ex b being Element of pne2 st ( a in ppns2 & $1 = [a,b] & $2 = a +* (n .--> b) ); A22: for x being Element of CP9 ex y being Element of CPP st S1[x,y] proof let x be Element of CP9; ::_thesis: ex y being Element of CPP st S1[x,y] reconsider xx = x as Element of [:ppns2,pne2:] ; set a = xx `1 ; set b = xx `2 ; reconsider a = xx `1 as Element of ppns2 by MCART_1:10; reconsider b = xx `2 as Element of pne2 by MCART_1:10; A23: not product (Carrier (p | ns)) is empty by A21, YELLOW_1:def_4; A24: a is Element of product (Carrier (p | ns)) by YELLOW_1:def_4; then A25: ex g being Function st ( a = g & dom g = dom (Carrier (p | ns)) & ( for i being set st i in dom (Carrier (p | ns)) holds g . i in (Carrier (p | ns)) . i ) ) by A23, CARD_3:def_5; reconsider a = a as Function by A24; set y = a +* (n .--> b); now__::_thesis:_ex_g1_being_Function_st_ (_a_+*_(n_.-->_b)_=_g1_&_dom_g1_=_dom_(Carrier_p)_&_(_for_x_being_set_st_x_in_dom_(Carrier_p)_holds_ g1_._x_in_(Carrier_p)_._x_)_) set g1 = a +* (n .--> b); reconsider g1 = a +* (n .--> b) as Function ; take g1 = g1; ::_thesis: ( a +* (n .--> b) = g1 & dom g1 = dom (Carrier p) & ( for x being set st x in dom (Carrier p) holds b2 . b3 in (Carrier p) . b3 ) ) thus a +* (n .--> b) = g1 ; ::_thesis: ( dom g1 = dom (Carrier p) & ( for x being set st x in dom (Carrier p) holds b2 . b3 in (Carrier p) . b3 ) ) A26: dom a = ns by A25, PARTFUN1:def_2; thus dom g1 = (dom a) \/ (dom (n .--> b)) by FUNCT_4:def_1 .= n \/ {n} by A1, A26, FUNCOP_1:13 .= n + 1 by AFINSQ_1:2 .= dom (Carrier p) by PARTFUN1:def_2 ; ::_thesis: for x being set st x in dom (Carrier p) holds b2 . b3 in (Carrier p) . b3 let x be set ; ::_thesis: ( x in dom (Carrier p) implies b1 . b2 in (Carrier p) . b2 ) assume x in dom (Carrier p) ; ::_thesis: b1 . b2 in (Carrier p) . b2 then A27: x in n + 1 ; then A28: x in n \/ {n} by AFINSQ_1:2; percases ( x in n or x in {n} ) by A28, XBOOLE_0:def_3; supposeA29: x in n ; ::_thesis: b1 . b2 in (Carrier p) . b2 then x <> n ; then not x in dom (n .--> b) by TARSKI:def_1; then A30: g1 . x = a . x by FUNCT_4:11; A31: x in dom (Carrier (p | ns)) by A1, A29, PARTFUN1:def_2; A32: ex R being 1-sorted st ( R = (p | ns) . x & (Carrier (p | ns)) . x = the carrier of R ) by A1, A29, PRALG_1:def_13; ex R2 being 1-sorted st ( R2 = p . x & (Carrier p) . x = the carrier of R2 ) by A27, PRALG_1:def_13; then (Carrier (p | ns)) . x = (Carrier p) . x by A1, A29, A32, FUNCT_1:49; hence g1 . x in (Carrier p) . x by A25, A30, A31; ::_thesis: verum end; supposeA33: x in {n} ; ::_thesis: b1 . b2 in (Carrier p) . b2 then A34: x = n by TARSKI:def_1; x in dom (n .--> b) by A33, FUNCOP_1:13; then A35: g1 . x = (n .--> b) . n by A34, FUNCT_4:13 .= b by FUNCOP_1:72 ; ex R being 1-sorted st ( R = p . ne & (Carrier p) . ne = the carrier of R ) by PRALG_1:def_13; hence g1 . x in (Carrier p) . x by A2, A34, A35; ::_thesis: verum end; end; end; then a +* (n .--> b) in product (Carrier p) by CARD_3:def_5; then reconsider y = a +* (n .--> b) as Element of CPP by YELLOW_1:def_4; take y ; ::_thesis: S1[x,y] take a ; ::_thesis: ex b being Element of pne2 st ( a in ppns2 & x = [a,b] & y = a +* (n .--> b) ) take b ; ::_thesis: ( a in ppns2 & x = [a,b] & y = a +* (n .--> b) ) thus a in ppns2 ; ::_thesis: ( x = [a,b] & y = a +* (n .--> b) ) thus x = [a,b] by MCART_1:21; ::_thesis: y = a +* (n .--> b) thus y = a +* (n .--> b) ; ::_thesis: verum end; consider f being Function of CP9,CPP such that A36: for x being Element of CP9 holds S1[x,f . x] from FUNCT_2:sch_3(A22); now__::_thesis:_ex_f_being_Function_of_[:(product_(p_|_ns)),(p_._ne):],(product_p)_st_f_is_isomorphic the carrier of [:(product (p | ns)),(p . ne):] = [: the carrier of (product (p | ns)), the carrier of (p . ne):] by YELLOW_3:def_2; then reconsider f = f as Function of [:(product (p | ns)),(p . ne):],(product p) ; reconsider f9 = f as Function of S,T ; take f = f; ::_thesis: f is isomorphic now__::_thesis:_for_x1,_x2_being_set_st_x1_in_dom_f_&_x2_in_dom_f_&_f_._x1_=_f_._x2_holds_ x1_=_x2 let x1, x2 be set ; ::_thesis: ( x1 in dom f & x2 in dom f & f . x1 = f . x2 implies x1 = x2 ) assume that A37: x1 in dom f and A38: x2 in dom f and A39: f . x1 = f . x2 ; ::_thesis: x1 = x2 x1 is Element of [: the carrier of (product (p | ns)), the carrier of (p . ne):] by A37, YELLOW_3:def_2; then consider a1 being Function, b1 being Element of pne2 such that A40: a1 in ppns2 and A41: x1 = [a1,b1] and A42: f . x1 = a1 +* (n .--> b1) by A36; x2 is Element of [: the carrier of (product (p | ns)), the carrier of (p . ne):] by A38, YELLOW_3:def_2; then consider a2 being Function, b2 being Element of pne2 such that A43: a2 in ppns2 and A44: x2 = [a2,b2] and A45: f . x2 = a2 +* (n .--> b2) by A36; a1 in product (Carrier (p | ns)) by A40, YELLOW_1:def_4; then A46: ex g1 being Function st ( a1 = g1 & dom g1 = dom (Carrier (p | ns)) & ( for x being set st x in dom (Carrier (p | ns)) holds g1 . x in (Carrier (p | ns)) . x ) ) by CARD_3:def_5; a2 in product (Carrier (p | ns)) by A43, YELLOW_1:def_4; then A47: ex g2 being Function st ( a2 = g2 & dom g2 = dom (Carrier (p | ns)) & ( for x being set st x in dom (Carrier (p | ns)) holds g2 . x in (Carrier (p | ns)) . x ) ) by CARD_3:def_5; A48: dom (n .--> b1) = {n} by FUNCOP_1:13; then A49: dom (n .--> b1) = dom (n .--> b2) by FUNCOP_1:13; A50: dom a1 = n by A1, A46, PARTFUN1:def_2; A51: now__::_thesis:_n_misses_{n} assume not n misses {n} ; ::_thesis: contradiction then n /\ {n} <> {} by XBOOLE_0:def_7; then consider x being set such that A52: x in n /\ {n} by XBOOLE_0:def_1; A53: x in n by A52, XBOOLE_0:def_4; x in {n} by A52, XBOOLE_0:def_4; then x = n by TARSKI:def_1; hence contradiction by A53; ::_thesis: verum end; then A54: dom a1 misses dom (n .--> b1) by A50, FUNCOP_1:13; A55: dom a2 misses dom (n .--> b2) by A46, A47, A50, A51, FUNCOP_1:13; A56: now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_a1_implies_[a,b]_in_a2_)_&_(_[a,b]_in_a2_implies_[a,b]_in_a1_)_) let a, b be set ; ::_thesis: ( ( [a,b] in a1 implies [a,b] in a2 ) & ( [a,b] in a2 implies [a,b] in a1 ) ) hereby ::_thesis: ( [a,b] in a2 implies [a,b] in a1 ) assume A57: [a,b] in a1 ; ::_thesis: [a,b] in a2 then A58: a in dom a1 by FUNCT_1:1; A59: b = a1 . a by A57, FUNCT_1:1; A60: a in (dom a1) \/ (dom (n .--> b1)) by A58, XBOOLE_0:def_3; then not a in dom (n .--> b1) by A54, A58, XBOOLE_0:5; then A61: (a2 +* (n .--> b2)) . a = a1 . a by A39, A42, A45, A60, FUNCT_4:def_1; A62: a in (dom a2) \/ (dom (n .--> b2)) by A46, A47, A58, XBOOLE_0:def_3; A63: a in dom a2 by A46, A47, A57, FUNCT_1:1; then not a in dom (n .--> b2) by A55, A62, XBOOLE_0:5; then b = a2 . a by A59, A61, A62, FUNCT_4:def_1; hence [a,b] in a2 by A63, FUNCT_1:1; ::_thesis: verum end; assume A64: [a,b] in a2 ; ::_thesis: [a,b] in a1 then A65: a in dom a2 by FUNCT_1:1; A66: b = a2 . a by A64, FUNCT_1:1; A67: a in (dom a2) \/ (dom (n .--> b2)) by A65, XBOOLE_0:def_3; then not a in dom (n .--> b2) by A55, A65, XBOOLE_0:5; then A68: (a1 +* (n .--> b1)) . a = a2 . a by A39, A42, A45, A67, FUNCT_4:def_1; A69: a in (dom a1) \/ (dom (n .--> b1)) by A46, A47, A65, XBOOLE_0:def_3; A70: a in dom a1 by A46, A47, A64, FUNCT_1:1; then not a in dom (n .--> b1) by A54, A69, XBOOLE_0:5; then b = a1 . a by A66, A68, A69, FUNCT_4:def_1; hence [a,b] in a1 by A70, FUNCT_1:1; ::_thesis: verum end; A71: n in dom (n .--> b1) by A48, TARSKI:def_1; then A72: n in (dom a1) \/ (dom (n .--> b1)) by XBOOLE_0:def_3; A73: n in dom (n .--> b2) by A48, A49, TARSKI:def_1; then n in (dom a2) \/ (dom (n .--> b2)) by XBOOLE_0:def_3; then (a1 +* (n .--> b1)) . n = (n .--> b2) . n by A39, A42, A45, A73, FUNCT_4:def_1 .= b2 by FUNCOP_1:72 ; then b2 = (n .--> b1) . n by A71, A72, FUNCT_4:def_1 .= b1 by FUNCOP_1:72 ; hence x1 = x2 by A41, A44, A56, RELAT_1:def_2; ::_thesis: verum end; then A74: f is one-to-one by FUNCT_1:def_4; now__::_thesis:_for_y_being_set_holds_ (_(_y_in_rng_f_implies_y_in_the_carrier_of_(product_p)_)_&_(_y_in_the_carrier_of_(product_p)_implies_y_in_rng_f_)_) let y be set ; ::_thesis: ( ( y in rng f implies y in the carrier of (product p) ) & ( y in the carrier of (product p) implies y in rng f ) ) thus ( y in rng f implies y in the carrier of (product p) ) ; ::_thesis: ( y in the carrier of (product p) implies y in rng f ) assume y in the carrier of (product p) ; ::_thesis: y in rng f then y in product (Carrier p) by YELLOW_1:def_4; then consider g being Function such that A75: y = g and A76: dom g = dom (Carrier p) and A77: for x being set st x in dom (Carrier p) holds g . x in (Carrier p) . x by CARD_3:def_5; A78: dom (Carrier p) = n + 1 by PARTFUN1:def_2; A79: n + 1 = { i where i is Element of NAT : i < n + 1 } by AXIOMS:4; A80: n in NAT by ORDINAL1:def_12; n < n + 1 by NAT_1:13; then A81: n in n + 1 by A79, A80; set a = g | n; set b = g . n; set x = [(g | n),(g . n)]; A82: dom (Carrier (p | ns)) = ns by PARTFUN1:def_2; A83: dom (g | n) = (dom g) /\ n by RELAT_1:61 .= (n + 1) /\ n by A76, PARTFUN1:def_2 ; then A84: dom (g | n) = n by Th2, XBOOLE_1:28; A85: dom (g | n) = dom (Carrier (p | ns)) by A1, A82, A83, XBOOLE_1:28; now__::_thesis:_for_x_being_set_st_x_in_dom_(Carrier_(p_|_ns))_holds_ (g_|_n)_._x_in_(Carrier_(p_|_ns))_._x let x be set ; ::_thesis: ( x in dom (Carrier (p | ns)) implies (g | n) . x in (Carrier (p | ns)) . x ) assume x in dom (Carrier (p | ns)) ; ::_thesis: (g | n) . x in (Carrier (p | ns)) . x then A86: x in n by A1; A87: n c= n + 1 by Th2; A88: (g | n) . x = g . x by A86, FUNCT_1:49; A89: g . x in (Carrier p) . x by A77, A78, A86, A87; A90: ex R1 being 1-sorted st ( R1 = p . x & (Carrier p) . x = the carrier of R1 ) by A86, A87, PRALG_1:def_13; ex R2 being 1-sorted st ( R2 = (p | ns) . x & (Carrier (p | ns)) . x = the carrier of R2 ) by A1, A86, PRALG_1:def_13; hence (g | n) . x in (Carrier (p | ns)) . x by A1, A86, A88, A89, A90, FUNCT_1:49; ::_thesis: verum end; then g | n in product (Carrier (p | ns)) by A85, CARD_3:9; then A91: g | n in the carrier of (product (p | ns)) by YELLOW_1:def_4; ex R1 being 1-sorted st ( R1 = p . n & (Carrier p) . n = the carrier of R1 ) by A81, PRALG_1:def_13; then A92: g . n in the carrier of (p . ne) by A2, A77, A78; then [(g | n),(g . n)] in [: the carrier of (product (p | ns)), the carrier of (p . ne):] by A91, ZFMISC_1:87; then A93: [(g | n),(g . n)] in dom f by FUNCT_2:def_1; [(g | n),(g . n)] is Element of CP9 by A91, A92, ZFMISC_1:87; then consider ta being Function, tb being Element of pne2 such that ta in ppns2 and A94: [(g | n),(g . n)] = [ta,tb] and A95: f . [(g | n),(g . n)] = ta +* (n .--> tb) by A36; A96: ta = g | n by A94, XTUPLE_0:1; A97: tb = g . n by A94, XTUPLE_0:1; now__::_thesis:_for_i,_j_being_set_holds_ (_(_[i,j]_in_(g_|_n)_+*_(n_.-->_(g_._n))_implies_[i,j]_in_g_)_&_(_[i,j]_in_g_implies_[i,j]_in_(g_|_n)_+*_(n_.-->_(g_._n))_)_) let i, j be set ; ::_thesis: ( ( [i,j] in (g | n) +* (n .--> (g . n)) implies [i,j] in g ) & ( [i,j] in g implies [b1,b2] in (g | n) +* (n .--> (g . n)) ) ) hereby ::_thesis: ( [i,j] in g implies [b1,b2] in (g | n) +* (n .--> (g . n)) ) assume A98: [i,j] in (g | n) +* (n .--> (g . n)) ; ::_thesis: [i,j] in g then i in dom ((g | n) +* (n .--> (g . n))) by FUNCT_1:1; then A99: i in (dom (g | n)) \/ (dom (n .--> (g . n))) by FUNCT_4:def_1; then A100: i in n \/ {n} by A84, FUNCOP_1:13; A101: ((g | n) +* (n .--> (g . n))) . i = j by A98, FUNCT_1:1; percases ( i in dom (n .--> (g . n)) or not i in dom (n .--> (g . n)) ) ; supposeA102: i in dom (n .--> (g . n)) ; ::_thesis: [i,j] in g then i in {n} ; then A103: i = n by TARSKI:def_1; ((g | n) +* (n .--> (g . n))) . i = (n .--> (g . n)) . i by A99, A102, FUNCT_4:def_1 .= g . n by A103, FUNCOP_1:72 ; then A104: g . i = j by A98, A103, FUNCT_1:1; i in dom g by A76, A78, A100, AFINSQ_1:2; hence [i,j] in g by A104, FUNCT_1:1; ::_thesis: verum end; supposeA105: not i in dom (n .--> (g . n)) ; ::_thesis: [i,j] in g then not i in {n} by FUNCOP_1:13; then A106: i in n by A100, XBOOLE_0:def_3; ((g | n) +* (n .--> (g . n))) . i = (g | n) . i by A99, A105, FUNCT_4:def_1; then A107: g . i = j by A101, A106, FUNCT_1:49; i in dom g by A76, A78, A100, AFINSQ_1:2; hence [i,j] in g by A107, FUNCT_1:1; ::_thesis: verum end; end; end; assume A108: [i,j] in g ; ::_thesis: [b1,b2] in (g | n) +* (n .--> (g . n)) then A109: i in n + 1 by A76, A78, FUNCT_1:1; then A110: i in n \/ {n} by AFINSQ_1:2; dom ((g | n) +* (n .--> (g . n))) = (dom (g | n)) \/ (dom (n .--> (g . n))) by FUNCT_4:def_1 .= n \/ {n} by A84, FUNCOP_1:13 ; then A111: i in dom ((g | n) +* (n .--> (g . n))) by A109, AFINSQ_1:2; then A112: i in (dom (g | n)) \/ (dom (n .--> (g . n))) by FUNCT_4:def_1; percases ( i in dom (n .--> (g . n)) or not i in dom (n .--> (g . n)) ) ; supposeA113: i in dom (n .--> (g . n)) ; ::_thesis: [b1,b2] in (g | n) +* (n .--> (g . n)) then i in {n} ; then A114: i = n by TARSKI:def_1; ((g | n) +* (n .--> (g . n))) . i = (n .--> (g . n)) . i by A112, A113, FUNCT_4:def_1 .= g . n by A114, FUNCOP_1:72 .= j by A108, A114, FUNCT_1:1 ; hence [i,j] in (g | n) +* (n .--> (g . n)) by A111, FUNCT_1:1; ::_thesis: verum end; supposeA115: not i in dom (n .--> (g . n)) ; ::_thesis: [b1,b2] in (g | n) +* (n .--> (g . n)) then not i in {n} by FUNCOP_1:13; then A116: i in n by A110, XBOOLE_0:def_3; ((g | n) +* (n .--> (g . n))) . i = (g | n) . i by A112, A115, FUNCT_4:def_1 .= g . i by A116, FUNCT_1:49 .= j by A108, FUNCT_1:1 ; hence [i,j] in (g | n) +* (n .--> (g . n)) by A111, FUNCT_1:1; ::_thesis: verum end; end; end; then f . [(g | n),(g . n)] = y by A75, A95, A96, A97, RELAT_1:def_2; hence y in rng f by A93, FUNCT_1:def_3; ::_thesis: verum end; then A117: rng f = the carrier of (product p) by TARSKI:1; now__::_thesis:_for_x,_y_being_Element_of_S_holds_ (_(_x_<=_y_implies_f9_._x_<=_f9_._y_)_&_(_f9_._x_<=_f9_._y_implies_x_<=_y_)_) let x, y be Element of S; ::_thesis: ( ( x <= y implies f9 . x <= f9 . y ) & ( f9 . x <= f9 . y implies x <= y ) ) A118: x is Element of [: the carrier of (product (p | ns)), the carrier of (p . ne):] by YELLOW_3:def_2; then consider xa being Function, xb being Element of pne2 such that A119: xa in ppns2 and A120: x = [xa,xb] and A121: f . x = xa +* (n .--> xb) by A36; dom f = CP9 by FUNCT_2:def_1; then f . x in the carrier of (product p) by A117, A118, FUNCT_1:def_3; then A122: f9 . x in product (Carrier p) by YELLOW_1:def_4; y is Element of [: the carrier of (product (p | ns)), the carrier of (p . ne):] by YELLOW_3:def_2; then consider ya being Function, yb being Element of pne2 such that A123: ya in ppns2 and A124: y = [ya,yb] and A125: f . y = ya +* (n .--> yb) by A36; A126: [xa,xb] `1 = xa ; A127: [xa,xb] `2 = xb ; A128: xa in product (Carrier (p | ns)) by A119, YELLOW_1:def_4; then A129: ex gx being Function st ( xa = gx & dom gx = dom (Carrier (p | ns)) & ( for x being set st x in dom (Carrier (p | ns)) holds gx . x in (Carrier (p | ns)) . x ) ) by CARD_3:def_5; then A130: dom xa = n by A1, PARTFUN1:def_2; then A131: (dom xa) \/ (dom (n .--> xb)) = n \/ {n} by FUNCOP_1:13; ya in product (Carrier (p | ns)) by A123, YELLOW_1:def_4; then A132: ex gy being Function st ( ya = gy & dom gy = dom (Carrier (p | ns)) & ( for x being set st x in dom (Carrier (p | ns)) holds gy . x in (Carrier (p | ns)) . x ) ) by CARD_3:def_5; then A133: dom ya = n by A1, PARTFUN1:def_2; then A134: (dom ya) \/ (dom (n .--> yb)) = n \/ {n} by FUNCOP_1:13; reconsider xa9 = xa as Element of (product (p | ns)) by A119; reconsider ya9 = ya as Element of (product (p | ns)) by A123; hereby ::_thesis: ( f9 . x <= f9 . y implies x <= y ) assume x <= y ; ::_thesis: f9 . x <= f9 . y then [x,y] in the InternalRel of S by ORDERS_2:def_5; then A135: [x,y] in [" the InternalRel of (product (p | ns)), the InternalRel of (p . ne)"] by YELLOW_3:def_2; then A136: [(([x,y] `1) `1),(([x,y] `2) `1)] in the InternalRel of (product (p | ns)) by YELLOW_3:10; A137: [(([x,y] `1) `2),(([x,y] `2) `2)] in the InternalRel of (p . ne) by A135, YELLOW_3:10; [ya,yb] `1 = ya ; then A138: [xa,ya] in the InternalRel of (product (p | ns)) by A124, A126, A136, A120; A139: xa in product (Carrier (p | ns)) by A119, YELLOW_1:def_4; xa9 <= ya9 by A138, ORDERS_2:def_5; then consider ffx, ggx being Function such that A140: ffx = xa and A141: ggx = ya and A142: for i being set st i in n holds ex RR being RelStr ex xxi, yyi being Element of RR st ( RR = (p | ns) . i & xxi = ffx . i & yyi = ggx . i & xxi <= yyi ) by A1, A139, YELLOW_1:def_4; [ya,yb] `2 = yb ; then A143: [xb,yb] in the InternalRel of (p . ne) by A124, A127, A137, A120; reconsider xb9 = xb as Element of (p . ne) ; reconsider yb9 = yb as Element of (p . ne) ; A144: xb9 <= yb9 by A143, ORDERS_2:def_5; ex f1, g1 being Function st ( f1 = f . x & g1 = f . y & ( for i being set st i in n + 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) ) ) proof set f1 = xa +* (n .--> xb); set g1 = ya +* (n .--> yb); take xa +* (n .--> xb) ; ::_thesis: ex g1 being Function st ( xa +* (n .--> xb) = f . x & g1 = f . y & ( for i being set st i in n + 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = g1 . i & xi <= yi ) ) ) take ya +* (n .--> yb) ; ::_thesis: ( xa +* (n .--> xb) = f . x & ya +* (n .--> yb) = f . y & ( for i being set st i in n + 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) ) ) thus xa +* (n .--> xb) = f . x by A121; ::_thesis: ( ya +* (n .--> yb) = f . y & ( for i being set st i in n + 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) ) ) thus ya +* (n .--> yb) = f . y by A125; ::_thesis: for i being set st i in n + 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) let i be set ; ::_thesis: ( i in n + 1 implies ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) ) assume A145: i in n + 1 ; ::_thesis: ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) A146: i in n \/ {n} by A145, AFINSQ_1:2; set R = p . i; set xi = (xa +* (n .--> xb)) . i; set yi = (ya +* (n .--> yb)) . i; i in dom p by A145, PARTFUN1:def_2; then p . i in rng p by FUNCT_1:def_3; then reconsider R = p . i as RelStr by YELLOW_1:def_3; A147: i in (dom xa) \/ (dom (n .--> xb)) by A131, A145, AFINSQ_1:2; now__::_thesis:_(xa_+*_(n_.-->_xb))_._i_is_Element_of_R percases ( i in dom (n .--> xb) or not i in dom (n .--> xb) ) ; supposeA148: i in dom (n .--> xb) ; ::_thesis: (xa +* (n .--> xb)) . i is Element of R then i in {n} ; then A149: i = n by TARSKI:def_1; (xa +* (n .--> xb)) . i = (n .--> xb) . i by A147, A148, FUNCT_4:def_1; hence (xa +* (n .--> xb)) . i is Element of R by A2, A149, FUNCOP_1:72; ::_thesis: verum end; supposeA150: not i in dom (n .--> xb) ; ::_thesis: (xa +* (n .--> xb)) . i is Element of R then A151: not i in {n} by FUNCOP_1:13; then A152: i in n by A146, XBOOLE_0:def_3; A153: i in dom (Carrier (p | ns)) by A3, A146, A151, XBOOLE_0:def_3; (xa +* (n .--> xb)) . i = xa . i by A147, A150, FUNCT_4:def_1; then A154: (xa +* (n .--> xb)) . i in (Carrier (p | ns)) . i by A129, A153; ex R2 being 1-sorted st ( R2 = (p | ns) . i & (Carrier (p | ns)) . i = the carrier of R2 ) by A1, A152, PRALG_1:def_13; hence (xa +* (n .--> xb)) . i is Element of R by A1, A152, A154, FUNCT_1:49; ::_thesis: verum end; end; end; then reconsider xi = (xa +* (n .--> xb)) . i as Element of R ; A155: i in (dom ya) \/ (dom (n .--> yb)) by A134, A145, AFINSQ_1:2; now__::_thesis:_(ya_+*_(n_.-->_yb))_._i_is_Element_of_R percases ( i in dom (n .--> yb) or not i in dom (n .--> yb) ) ; supposeA156: i in dom (n .--> yb) ; ::_thesis: (ya +* (n .--> yb)) . i is Element of R then i in {n} ; then A157: i = n by TARSKI:def_1; (ya +* (n .--> yb)) . i = (n .--> yb) . i by A155, A156, FUNCT_4:def_1; hence (ya +* (n .--> yb)) . i is Element of R by A2, A157, FUNCOP_1:72; ::_thesis: verum end; supposeA158: not i in dom (n .--> yb) ; ::_thesis: (ya +* (n .--> yb)) . i is Element of R then A159: not i in {n} by FUNCOP_1:13; then A160: i in n by A146, XBOOLE_0:def_3; A161: i in dom (Carrier (p | ns)) by A3, A146, A159, XBOOLE_0:def_3; (ya +* (n .--> yb)) . i = ya . i by A155, A158, FUNCT_4:def_1; then A162: (ya +* (n .--> yb)) . i in (Carrier (p | ns)) . i by A132, A161; ex R2 being 1-sorted st ( R2 = (p | ns) . i & (Carrier (p | ns)) . i = the carrier of R2 ) by A1, A160, PRALG_1:def_13; hence (ya +* (n .--> yb)) . i is Element of R by A1, A160, A162, FUNCT_1:49; ::_thesis: verum end; end; end; then reconsider yi = (ya +* (n .--> yb)) . i as Element of R ; take R ; ::_thesis: ex xi, yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) take xi ; ::_thesis: ex yi being Element of R st ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) take yi ; ::_thesis: ( R = p . i & xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) thus R = p . i ; ::_thesis: ( xi = (xa +* (n .--> xb)) . i & yi = (ya +* (n .--> yb)) . i & xi <= yi ) thus xi = (xa +* (n .--> xb)) . i ; ::_thesis: ( yi = (ya +* (n .--> yb)) . i & xi <= yi ) thus yi = (ya +* (n .--> yb)) . i ; ::_thesis: xi <= yi percases ( i in n or i in {n} ) by A146, XBOOLE_0:def_3; supposeA163: i in n ; ::_thesis: xi <= yi A164: not i in {n} proof assume i in {n} ; ::_thesis: contradiction then n in n by A163, TARSKI:def_1; hence contradiction ; ::_thesis: verum end; then A165: not i in dom (n .--> xb) ; not i in dom (n .--> yb) by A164; then A166: yi = ya . i by A155, FUNCT_4:def_1; consider RR being RelStr , xxi, yyi being Element of RR such that A167: RR = (p | ns) . i and A168: xxi = ffx . i and A169: yyi = ggx . i and A170: xxi <= yyi by A142, A163; RR = R by A1, A163, A167, FUNCT_1:49; hence xi <= yi by A140, A141, A147, A165, A166, A168, A169, A170, FUNCT_4:def_1; ::_thesis: verum end; supposeA171: i in {n} ; ::_thesis: xi <= yi then A172: i = n by TARSKI:def_1; A173: i in dom (n .--> xb) by A171, FUNCOP_1:13; A174: i in dom (n .--> yb) by A171, FUNCOP_1:13; A175: xi = (n .--> xb) . i by A147, A173, FUNCT_4:def_1 .= xb by A172, FUNCOP_1:72 ; yi = (n .--> yb) . i by A155, A174, FUNCT_4:def_1 .= yb by A172, FUNCOP_1:72 ; hence xi <= yi by A2, A144, A171, A175, TARSKI:def_1; ::_thesis: verum end; end; end; hence f9 . x <= f9 . y by A122, YELLOW_1:def_4; ::_thesis: verum end; assume f9 . x <= f9 . y ; ::_thesis: x <= y then consider f1, g1 being Function such that A176: f1 = f . x and A177: g1 = f . y and A178: for i being set st i in n + 1 holds ex R being RelStr ex xi, yi being Element of R st ( R = p . i & xi = f1 . i & yi = g1 . i & xi <= yi ) by A122, YELLOW_1:def_4; now__::_thesis:_ex_f2,_g2_being_Function_st_ (_f2_=_xa9_&_g2_=_ya9_&_(_for_i_being_set_st_i_in_ns_holds_ ex_R_being_RelStr_ex_xi,_yi_being_Element_of_R_st_ (_R_=_(p_|_ns)_._i_&_xi_=_f2_._i_&_yi_=_g2_._i_&_xi_<=_yi_)_)_) set f2 = xa9; set g2 = ya9; reconsider f2 = xa9, g2 = ya9 as Function ; take f2 = f2; ::_thesis: ex g2 being Function st ( f2 = xa9 & g2 = ya9 & ( for i being set st i in ns holds ex R being RelStr ex xi, yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) ) ) take g2 = g2; ::_thesis: ( f2 = xa9 & g2 = ya9 & ( for i being set st i in ns holds ex R being RelStr ex xi, yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) ) ) thus ( f2 = xa9 & g2 = ya9 ) ; ::_thesis: for i being set st i in ns holds ex R being RelStr ex xi, yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) let i be set ; ::_thesis: ( i in ns implies ex R being RelStr ex xi, yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) ) assume A179: i in ns ; ::_thesis: ex R being RelStr ex xi, yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) A180: not i in {n} proof assume i in {n} ; ::_thesis: contradiction then i = n by TARSKI:def_1; hence contradiction by A1, A179; ::_thesis: verum end; then A181: not i in dom (n .--> yb) ; A182: not i in dom (n .--> xb) by A180; set R = (p | ns) . i; i in dom (p | ns) by A179, PARTFUN1:def_2; then (p | ns) . i in rng (p | ns) by FUNCT_1:def_3; then reconsider R = (p | ns) . i as RelStr by YELLOW_1:def_3; take R = R; ::_thesis: ex xi, yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) set xi = xa . i; set yi = ya . i; A183: i in (dom xa) \/ (dom (n .--> xb)) by A1, A130, A179, XBOOLE_0:def_3; A184: i in dom (Carrier (p | ns)) by A179, PARTFUN1:def_2; ex R2 being 1-sorted st ( R2 = (p | ns) . i & (Carrier (p | ns)) . i = the carrier of R2 ) by A179, PRALG_1:def_13; then reconsider xi = xa . i as Element of R by A129, A184; A185: i in (dom ya) \/ (dom (n .--> yb)) by A1, A133, A179, XBOOLE_0:def_3; ex R2 being 1-sorted st ( R2 = (p | ns) . i & (Carrier (p | ns)) . i = the carrier of R2 ) by A179, PRALG_1:def_13; then reconsider yi = ya . i as Element of R by A132, A184; take xi = xi; ::_thesis: ex yi being Element of R st ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) take yi = yi; ::_thesis: ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i & xi <= yi ) thus ( R = (p | ns) . i & xi = f2 . i & yi = g2 . i ) ; ::_thesis: xi <= yi consider R2 being RelStr , xi2, yi2 being Element of R2 such that A186: R2 = p . i and A187: xi2 = f1 . i and A188: yi2 = g1 . i and A189: xi2 <= yi2 by A178, A179; A190: R2 = R by A179, A186, FUNCT_1:49; (xa +* (n .--> xb)) . i = xi by A182, A183, FUNCT_4:def_1; hence xi <= yi by A121, A125, A176, A177, A181, A185, A187, A188, A189, A190, FUNCT_4:def_1; ::_thesis: verum end; then xa9 <= ya9 by A128, YELLOW_1:def_4; then A191: [xa,ya] in the InternalRel of (product (p | ns)) by ORDERS_2:def_5; [ya,yb] `1 = ya ; then A192: [(([x,y] `1) `1),(([x,y] `2) `1)] in the InternalRel of (product (p | ns)) by A124, A126, A120, A191; consider Rn being RelStr , xni, yni being Element of Rn such that A193: Rn = p . ne and A194: xni = f1 . n and A195: yni = g1 . n and A196: xni <= yni by A2, A178; A197: [xni,yni] in the InternalRel of (p . ne) by A193, A196, ORDERS_2:def_5; dom (n .--> xb) = {n} by FUNCOP_1:13; then A198: n in dom (n .--> xb) by TARSKI:def_1; then n in (dom xa) \/ (dom (n .--> xb)) by XBOOLE_0:def_3; then A199: (xa +* (n .--> xb)) . n = (n .--> xb) . n by A198, FUNCT_4:def_1 .= xb by FUNCOP_1:72 ; dom (n .--> yb) = {n} by FUNCOP_1:13; then A200: n in dom (n .--> yb) by TARSKI:def_1; then n in (dom ya) \/ (dom (n .--> yb)) by XBOOLE_0:def_3; then A201: (ya +* (n .--> yb)) . n = (n .--> yb) . n by A200, FUNCT_4:def_1 .= yb by FUNCOP_1:72 ; [ya,yb] `2 = yb ; then A202: [(([x,y] `1) `2),(([x,y] `2) `2)] in the InternalRel of (p . ne) by A121, A124, A125, A127, A176, A177, A194, A195, A197, A199, A201, A120; A203: [x,y] `1 = [xa,xb] by A120; [x,y] `2 = [ya,yb] by A124; then [x,y] in [" the InternalRel of (product (p | ns)), the InternalRel of (p . ne)"] by A192, A202, A203, YELLOW_3:10; then [x,y] in the InternalRel of S by YELLOW_3:def_2; hence x <= y by ORDERS_2:def_5; ::_thesis: verum end; hence f is isomorphic by A74, A117, WAYBEL_0:66; ::_thesis: verum end; hence [:(product (p | ns)),(p . ne):], product p are_isomorphic by WAYBEL_1:def_8; ::_thesis: verum end; end; end; theorem Th42: :: DICKSON:42 for n being non empty Nat for p being RelStr-yielding ManySortedSet of n st ( for i being Element of n holds ( p . i is Dickson & p . i is quasi_ordered ) ) holds ( product p is quasi_ordered & product p is Dickson ) proof defpred S1[ non empty Nat] means for p being RelStr-yielding ManySortedSet of $1 st ( for i being Element of $1 holds ( p . i is Dickson & p . i is quasi_ordered ) ) holds ( product p is quasi_ordered & product p is Dickson ); A1: S1[1] proof let p be RelStr-yielding ManySortedSet of 1; ::_thesis: ( ( for i being Element of 1 holds ( p . i is Dickson & p . i is quasi_ordered ) ) implies ( product p is quasi_ordered & product p is Dickson ) ) assume A2: for i being Element of 1 holds ( p . i is Dickson & p . i is quasi_ordered ) ; ::_thesis: ( product p is quasi_ordered & product p is Dickson ) reconsider z = 0 as Element of 1 by CARD_1:49, TARSKI:def_1; A3: p . z is Dickson by A2; A4: p . z is quasi_ordered by A2; p . z, product p are_isomorphic by Th39; hence ( product p is quasi_ordered & product p is Dickson ) by A3, A4, Th38; ::_thesis: verum end; A5: now__::_thesis:_for_n_being_non_empty_Nat_st_S1[n]_holds_ S1[n_+_1] let n be non empty Nat; ::_thesis: ( S1[n] implies S1[n + 1] ) assume A6: S1[n] ; ::_thesis: S1[n + 1] thus S1[n + 1] ::_thesis: verum proof let p be RelStr-yielding ManySortedSet of n + 1; ::_thesis: ( ( for i being Element of n + 1 holds ( p . i is Dickson & p . i is quasi_ordered ) ) implies ( product p is quasi_ordered & product p is Dickson ) ) assume A7: for i being Element of n + 1 holds ( p . i is Dickson & p . i is quasi_ordered ) ; ::_thesis: ( product p is quasi_ordered & product p is Dickson ) n <= n + 1 by NAT_1:11; then reconsider ns = n as Subset of (n + 1) by NAT_1:39; A8: n + 1 = { k where k is Element of NAT : k < n + 1 } by AXIOMS:4; A9: n in NAT by ORDINAL1:def_12; n < n + 1 by NAT_1:13; then n in n + 1 by A8, A9; then reconsider ne = n as Element of n + 1 ; reconsider pns = p | ns as RelStr-yielding ManySortedSet of n ; A10: for i being Element of n holds ( pns . i is Dickson & pns . i is quasi_ordered ) proof let i be Element of n; ::_thesis: ( pns . i is Dickson & pns . i is quasi_ordered ) A11: pns . i = p . i by FUNCT_1:49; A12: n = { k where k is Element of NAT : k < n } by AXIOMS:4; i in n ; then consider k being Element of NAT such that A13: k = i and A14: k < n by A12; k < n + 1 by A14, NAT_1:13; then i in n + 1 by A8, A13; then reconsider i9 = i as Element of n + 1 ; i9 = i ; hence ( pns . i is Dickson & pns . i is quasi_ordered ) by A7, A11; ::_thesis: verum end; then A15: product pns is Dickson by A6; A16: product pns is quasi_ordered by A6, A10; A17: p . ne is Dickson by A7; A18: p . ne is quasi_ordered by A7; then A19: [:(product (p | ns)),(p . ne):] is Dickson by A15, A16, A17, Th37; A20: [:(product (p | ns)),(p . ne):] is quasi_ordered by A15, A16, A17, A18, Th37; [:(product (p | ns)),(p . ne):], product p are_isomorphic by Th41; hence ( product p is quasi_ordered & product p is Dickson ) by A19, A20, Th38; ::_thesis: verum end; end; thus for n being non empty Nat holds S1[n] from NAT_1:sch_10(A1, A5); ::_thesis: verum end; Lm1: for p being RelStr-yielding ManySortedSet of {} holds ( not product p is empty & product p is quasi_ordered & product p is Dickson & product p is antisymmetric ) proof let p be RelStr-yielding ManySortedSet of {} ; ::_thesis: ( not product p is empty & product p is quasi_ordered & product p is Dickson & product p is antisymmetric ) A1: product p = RelStr(# {{}},(id {{}}) #) by YELLOW_1:26; set pp = product p; set cpp = the carrier of (product p); set ipp = the InternalRel of (product p); A2: the InternalRel of (product p) = id {{}} by YELLOW_1:26; thus not product p is empty by YELLOW_1:26; ::_thesis: ( product p is quasi_ordered & product p is Dickson & product p is antisymmetric ) thus product p is quasi_ordered ::_thesis: ( product p is Dickson & product p is antisymmetric ) proof thus product p is reflexive by YELLOW_1:26; :: according to DICKSON:def_3 ::_thesis: product p is transitive let x, y, z be set ; :: according to RELAT_2:def_8,ORDERS_2:def_3 ::_thesis: ( not x in the carrier of (product p) or not y in the carrier of (product p) or not z in the carrier of (product p) or not [x,y] in the InternalRel of (product p) or not [y,z] in the InternalRel of (product p) or [x,z] in the InternalRel of (product p) ) assume that x in the carrier of (product p) and y in the carrier of (product p) and z in the carrier of (product p) and A3: [x,y] in the InternalRel of (product p) and A4: [y,z] in the InternalRel of (product p) ; ::_thesis: [x,z] in the InternalRel of (product p) thus [x,z] in the InternalRel of (product p) by A2, A3, A4, RELAT_1:def_10; ::_thesis: verum end; thus product p is Dickson ::_thesis: product p is antisymmetric proof let N be Subset of the carrier of (product p); :: according to DICKSON:def_10 ::_thesis: ex B being set st ( B is_Dickson-basis_of N, product p & B is finite ) percases ( N = {} or N = {{}} ) by A1, ZFMISC_1:33; supposeA5: N = {} ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N, product p & B is finite ) take B = {} ; ::_thesis: ( B is_Dickson-basis_of N, product p & B is finite ) N = {} the carrier of (product p) by A5; hence B is_Dickson-basis_of N, product p by Th25; ::_thesis: B is finite thus B is finite ; ::_thesis: verum end; supposeA6: N = {{}} ; ::_thesis: ex B being set st ( B is_Dickson-basis_of N, product p & B is finite ) take B = {{}}; ::_thesis: ( B is_Dickson-basis_of N, product p & B is finite ) thus B is_Dickson-basis_of N, product p ::_thesis: B is finite proof thus B c= N by A6; :: according to DICKSON:def_9 ::_thesis: for a being Element of (product p) st a in N holds ex b being Element of (product p) st ( b in B & b <= a ) let a be Element of (product p); ::_thesis: ( a in N implies ex b being Element of (product p) st ( b in B & b <= a ) ) assume A7: a in N ; ::_thesis: ex b being Element of (product p) st ( b in B & b <= a ) take b = a; ::_thesis: ( b in B & b <= a ) thus b in B by A6, A7; ::_thesis: b <= a [b,a] in id {{}} by A6, A7, RELAT_1:def_10; hence b <= a by A2, ORDERS_2:def_5; ::_thesis: verum end; thus B is finite ; ::_thesis: verum end; end; end; thus product p is antisymmetric by YELLOW_1:26; ::_thesis: verum end; registration let p be RelStr-yielding ManySortedSet of {} ; cluster product p -> non empty ; coherence not product p is empty by Lm1; cluster product p -> antisymmetric ; coherence product p is antisymmetric by Lm1; cluster product p -> quasi_ordered ; coherence product p is quasi_ordered by Lm1; cluster product p -> Dickson ; coherence product p is Dickson by Lm1; end; definition func NATOrd -> Relation of NAT equals :: DICKSON:def 14 { [x,y] where x, y is Element of NAT : x <= y } ; correctness coherence { [x,y] where x, y is Element of NAT : x <= y } is Relation of NAT; proof set NO = { [x,y] where x, y is Element of NAT : x <= y } ; now__::_thesis:_for_z_being_set_st_z_in__{__[x,y]_where_x,_y_is_Element_of_NAT_:_x_<=_y__}__holds_ z_in_[:NAT,NAT:] let z be set ; ::_thesis: ( z in { [x,y] where x, y is Element of NAT : x <= y } implies z in [:NAT,NAT:] ) assume z in { [x,y] where x, y is Element of NAT : x <= y } ; ::_thesis: z in [:NAT,NAT:] then ex x, y being Element of NAT st ( z = [x,y] & x <= y ) ; hence z in [:NAT,NAT:] ; ::_thesis: verum end; hence { [x,y] where x, y is Element of NAT : x <= y } is Relation of NAT by TARSKI:def_3; ::_thesis: verum end; end; :: deftheorem defines NATOrd DICKSON:def_14_:_ NATOrd = { [x,y] where x, y is Element of NAT : x <= y } ; theorem Th43: :: DICKSON:43 NATOrd is_reflexive_in NAT proof let x be set ; :: according to RELAT_2:def_1 ::_thesis: ( not x in NAT or [x,x] in NATOrd ) assume x in NAT ; ::_thesis: [x,x] in NATOrd then reconsider x9 = x as Element of NAT ; x9 <= x9 ; hence [x,x] in NATOrd ; ::_thesis: verum end; theorem Th44: :: DICKSON:44 NATOrd is_antisymmetric_in NAT proof let x, y be set ; :: according to RELAT_2:def_4 ::_thesis: ( not x in NAT or not y in NAT or not [x,y] in NATOrd or not [y,x] in NATOrd or x = y ) assume that x in NAT and y in NAT and A1: [x,y] in NATOrd and A2: [y,x] in NATOrd ; ::_thesis: x = y consider x9, y9 being Element of NAT such that A3: [x,y] = [x9,y9] and A4: x9 <= y9 by A1; A5: x = x9 by A3, XTUPLE_0:1; A6: y = y9 by A3, XTUPLE_0:1; consider y2, x2 being Element of NAT such that A7: [y,x] = [y2,x2] and A8: y2 <= x2 by A2; A9: y2 = y9 by A6, A7, XTUPLE_0:1; x2 = x9 by A5, A7, XTUPLE_0:1; hence x = y by A4, A5, A6, A8, A9, XXREAL_0:1; ::_thesis: verum end; theorem Th45: :: DICKSON:45 NATOrd is_strongly_connected_in NAT proof now__::_thesis:_for_x,_y_being_set_st_x_in_NAT_&_y_in_NAT_&_not_[x,y]_in_NATOrd_holds_ [y,x]_in_NATOrd let x, y be set ; ::_thesis: ( x in NAT & y in NAT & not [x,y] in NATOrd implies [y,x] in NATOrd ) assume that A1: x in NAT and A2: y in NAT ; ::_thesis: ( [x,y] in NATOrd or [y,x] in NATOrd ) thus ( [x,y] in NATOrd or [y,x] in NATOrd ) ::_thesis: verum proof assume that A3: not [x,y] in NATOrd and A4: not [y,x] in NATOrd ; ::_thesis: contradiction reconsider x = x, y = y as Element of NAT by A1, A2; percases ( x <= y or y <= x ) ; suppose x <= y ; ::_thesis: contradiction hence contradiction by A3; ::_thesis: verum end; suppose y <= x ; ::_thesis: contradiction hence contradiction by A4; ::_thesis: verum end; end; end; end; hence NATOrd is_strongly_connected_in NAT by RELAT_2:def_7; ::_thesis: verum end; theorem Th46: :: DICKSON:46 NATOrd is_transitive_in NAT proof let x, y, z be set ; :: according to RELAT_2:def_8 ::_thesis: ( not x in NAT or not y in NAT or not z in NAT or not [x,y] in NATOrd or not [y,z] in NATOrd or [x,z] in NATOrd ) assume that x in NAT and y in NAT and z in NAT and A1: [x,y] in NATOrd and A2: [y,z] in NATOrd ; ::_thesis: [x,z] in NATOrd consider x1, y1 being Element of NAT such that A3: [x,y] = [x1,y1] and A4: x1 <= y1 by A1; A5: x = x1 by A3, XTUPLE_0:1; A6: y = y1 by A3, XTUPLE_0:1; consider y2, z2 being Element of NAT such that A7: [y,z] = [y2,z2] and A8: y2 <= z2 by A2; A9: y = y2 by A7, XTUPLE_0:1; A10: z = z2 by A7, XTUPLE_0:1; x1 <= z2 by A4, A6, A8, A9, XXREAL_0:2; hence [x,z] in NATOrd by A5, A10; ::_thesis: verum end; definition func OrderedNAT -> non empty RelStr equals :: DICKSON:def 15 RelStr(# NAT,NATOrd #); correctness coherence RelStr(# NAT,NATOrd #) is non empty RelStr ; ; end; :: deftheorem defines OrderedNAT DICKSON:def_15_:_ OrderedNAT = RelStr(# NAT,NATOrd #); Lm2: now__::_thesis:_OrderedNAT_is_connected now__::_thesis:_for_x,_y_being_Element_of_OrderedNAT_st_not_x_<=_y_holds_ y_<=_x let x, y be Element of OrderedNAT; ::_thesis: ( not x <= y implies y <= x ) assume not x <= y ; ::_thesis: y <= x then not [x,y] in NATOrd by ORDERS_2:def_5; then [y,x] in NATOrd by Th45, RELAT_2:def_7; hence y <= x by ORDERS_2:def_5; ::_thesis: verum end; hence OrderedNAT is connected by WAYBEL_0:def_29; ::_thesis: verum end; registration cluster OrderedNAT -> non empty connected ; coherence OrderedNAT is connected by Lm2; cluster OrderedNAT -> non empty Dickson ; coherence OrderedNAT is Dickson proof set IR9 = the InternalRel of (OrderedNAT \~); set CR9 = the carrier of (OrderedNAT \~); now__::_thesis:_for_N_being_set_st_N_c=_the_carrier_of_(OrderedNAT_\~)_&_N_<>_{}_holds_ ex_m_being_set_st_ (_m_in_N_&_the_InternalRel_of_(OrderedNAT_\~)_-Seg_m_misses_N_) let N be set ; ::_thesis: ( N c= the carrier of (OrderedNAT \~) & N <> {} implies ex m being set st ( m in N & the InternalRel of (OrderedNAT \~) -Seg m misses N ) ) assume that A1: N c= the carrier of (OrderedNAT \~) and A2: N <> {} ; ::_thesis: ex m being set st ( m in N & the InternalRel of (OrderedNAT \~) -Seg m misses N ) A3: ex k being set st k in N by A2, XBOOLE_0:def_1; defpred S1[ Nat] means c1 in N; A4: ex k being Nat st S1[k] by A1, A3; consider m being Nat such that A5: S1[m] and A6: for n being Nat st S1[n] holds m <= n from NAT_1:sch_5(A4); reconsider m = m as Element of OrderedNAT by ORDINAL1:def_12; thus ex m being set st ( m in N & the InternalRel of (OrderedNAT \~) -Seg m misses N ) ::_thesis: verum proof take m ; ::_thesis: ( m in N & the InternalRel of (OrderedNAT \~) -Seg m misses N ) thus m in N by A5; ::_thesis: the InternalRel of (OrderedNAT \~) -Seg m misses N now__::_thesis:_not_(_the_InternalRel_of_(OrderedNAT_\~)_-Seg_m)_/\_N_<>_{} assume ( the InternalRel of (OrderedNAT \~) -Seg m) /\ N <> {} ; ::_thesis: contradiction then consider z being set such that A7: z in ( the InternalRel of (OrderedNAT \~) -Seg m) /\ N by XBOOLE_0:def_1; A8: z in the InternalRel of (OrderedNAT \~) -Seg m by A7, XBOOLE_0:def_4; A9: z in N by A7, XBOOLE_0:def_4; A10: z <> m by A8, WELLORD1:1; A11: [z,m] in the InternalRel of (OrderedNAT \~) by A8, WELLORD1:1; then [z,m] in NATOrd ; then consider x, y being Element of NAT such that A12: [z,m] = [x,y] and x <= y ; A13: z = x by A12, XTUPLE_0:1; A14: m = y by A12, XTUPLE_0:1; then y <= x by A6, A9, A13; then [y,x] in NATOrd ; hence contradiction by A10, A11, A13, A14, Th44, RELAT_2:def_4; ::_thesis: verum end; hence the InternalRel of (OrderedNAT \~) -Seg m misses N by XBOOLE_0:def_7; ::_thesis: verum end; end; then the InternalRel of (OrderedNAT \~) is_well_founded_in the carrier of (OrderedNAT \~) by WELLORD1:def_3; then OrderedNAT \~ is well_founded by WELLFND1:def_2; hence OrderedNAT is Dickson by Th27; ::_thesis: verum end; cluster OrderedNAT -> non empty quasi_ordered ; coherence OrderedNAT is quasi_ordered proof A15: OrderedNAT is reflexive by Th43, ORDERS_2:def_2; OrderedNAT is transitive by Th46, ORDERS_2:def_3; hence OrderedNAT is quasi_ordered by A15, Def3; ::_thesis: verum end; cluster OrderedNAT -> non empty antisymmetric ; coherence OrderedNAT is antisymmetric by Th44, ORDERS_2:def_4; cluster OrderedNAT -> non empty transitive ; coherence OrderedNAT is transitive by Th46, ORDERS_2:def_3; cluster OrderedNAT -> non empty well_founded ; coherence OrderedNAT is well_founded proof set ir = the InternalRel of OrderedNAT; now__::_thesis:_for_f_being_sequence_of_OrderedNAT_holds_not_f_is_descending given f being sequence of OrderedNAT such that A16: f is descending ; ::_thesis: contradiction A17: dom f = NAT by FUNCT_2:def_1; reconsider rf = rng f as non empty Subset of NAT ; set m = min rf; min rf in rng f by XXREAL_2:def_7; then consider d being set such that A18: d in dom f and A19: min rf = f . d by FUNCT_1:def_3; reconsider d = d as Element of NAT by A18; A20: f . (d + 1) <> f . d by A16, WELLFND1:def_6; [(f . (d + 1)),(f . d)] in the InternalRel of OrderedNAT by A16, WELLFND1:def_6; then consider x, y being Element of NAT such that A21: [(f . (d + 1)),(f . d)] = [x,y] and A22: x <= y ; reconsider fd1 = f . (d + 1), fd = f . d as Element of NAT ; A23: f . (d + 1) = x by A21, XTUPLE_0:1; f . d = y by A21, XTUPLE_0:1; then A24: fd1 < fd by A20, A22, A23, XXREAL_0:1; f . (d + 1) in rf by A17, FUNCT_1:3; hence contradiction by A19, A24, XXREAL_2:def_7; ::_thesis: verum end; hence OrderedNAT is well_founded by WELLFND1:14; ::_thesis: verum end; end; Lm3: now__::_thesis:_for_n_being_Element_of_NAT_holds_ (_not_product_(n_-->_OrderedNAT)_is_empty_&_product_(n_-->_OrderedNAT)_is_Dickson_&_product_(n_-->_OrderedNAT)_is_quasi_ordered_&_product_(n_-->_OrderedNAT)_is_antisymmetric_) let n be Element of NAT ; ::_thesis: ( not product (b1 --> OrderedNAT) is empty & product (b1 --> OrderedNAT) is Dickson & product (b1 --> OrderedNAT) is quasi_ordered & product (b1 --> OrderedNAT) is antisymmetric ) set pp = product (n --> OrderedNAT); percases ( n is empty or not n is empty ) ; suppose n is empty ; ::_thesis: ( not product (b1 --> OrderedNAT) is empty & product (b1 --> OrderedNAT) is Dickson & product (b1 --> OrderedNAT) is quasi_ordered & product (b1 --> OrderedNAT) is antisymmetric ) hence ( not product (n --> OrderedNAT) is empty & product (n --> OrderedNAT) is Dickson & product (n --> OrderedNAT) is quasi_ordered & product (n --> OrderedNAT) is antisymmetric ) ; ::_thesis: verum end; suppose not n is empty ; ::_thesis: ( not product (b1 --> OrderedNAT) is empty & product (b1 --> OrderedNAT) is Dickson & product (b1 --> OrderedNAT) is quasi_ordered & product (b1 --> OrderedNAT) is antisymmetric ) then reconsider n9 = n as non empty Element of NAT ; set p = n9 --> OrderedNAT; thus not product (n --> OrderedNAT) is empty ; ::_thesis: ( product (n --> OrderedNAT) is Dickson & product (n --> OrderedNAT) is quasi_ordered & product (n --> OrderedNAT) is antisymmetric ) for i being Element of n9 holds ( (n9 --> OrderedNAT) . i is Dickson & (n9 --> OrderedNAT) . i is quasi_ordered ) by FUNCOP_1:7; hence ( product (n --> OrderedNAT) is Dickson & product (n --> OrderedNAT) is quasi_ordered ) by Th42; ::_thesis: product (n --> OrderedNAT) is antisymmetric product (n9 --> OrderedNAT) is antisymmetric ; hence product (n --> OrderedNAT) is antisymmetric ; ::_thesis: verum end; end; end; registration let n be Element of NAT ; cluster product (n --> OrderedNAT) -> non empty ; coherence not product (n --> OrderedNAT) is empty ; cluster product (n --> OrderedNAT) -> Dickson ; coherence product (n --> OrderedNAT) is Dickson by Lm3; cluster product (n --> OrderedNAT) -> quasi_ordered ; coherence product (n --> OrderedNAT) is quasi_ordered by Lm3; cluster product (n --> OrderedNAT) -> antisymmetric ; coherence product (n --> OrderedNAT) is antisymmetric by Lm3; end; theorem :: DICKSON:47 for M being RelStr st M is Dickson & M is quasi_ordered holds ( [:M,OrderedNAT:] is quasi_ordered & [:M,OrderedNAT:] is Dickson ) by Th37; theorem Th48: :: DICKSON:48 for R, S being non empty RelStr st R is Dickson & S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded proof let R, S be non empty RelStr ; ::_thesis: ( R is Dickson & S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S implies S \~ is well_founded ) assume that A1: R is Dickson and A2: S is quasi_ordered and A3: the InternalRel of R c= the InternalRel of S and A4: the carrier of R = the carrier of S ; ::_thesis: S \~ is well_founded S is Dickson by A1, A3, A4, Th28; hence S \~ is well_founded by A2, Th33; ::_thesis: verum end; theorem :: DICKSON:49 for R being non empty RelStr st R is quasi_ordered holds ( R is Dickson iff for S being non empty RelStr st S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded ) proof let R be non empty RelStr ; ::_thesis: ( R is quasi_ordered implies ( R is Dickson iff for S being non empty RelStr st S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded ) ) assume A1: R is quasi_ordered ; ::_thesis: ( R is Dickson iff for S being non empty RelStr st S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded ) A2: R is reflexive by A1, Def3; A3: R is transitive by A1, Def3; set IR = the InternalRel of R; set CR = the carrier of R; thus ( R is Dickson implies for S being non empty RelStr st S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded ) by Th48; ::_thesis: ( ( for S being non empty RelStr st S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded ) implies R is Dickson ) assume A4: for S being non empty RelStr st S is quasi_ordered & the InternalRel of R c= the InternalRel of S & the carrier of R = the carrier of S holds S \~ is well_founded ; ::_thesis: R is Dickson now__::_thesis:_R_is_Dickson assume not R is Dickson ; ::_thesis: contradiction then not for N being non empty Subset of R holds ( min-classes N is finite & not min-classes N is empty ) by A1, Th32; then consider f being sequence of R such that A5: for i, j being Element of NAT st i < j holds not f . i <= f . j by A1, Th31; defpred S1[ set , set ] means ( [$1,$2] in the InternalRel of R or ex i, j being Element of NAT st ( i < j & [$1,(f . j)] in the InternalRel of R & [(f . i),$2] in the InternalRel of R ) ); consider QOE being Relation of the carrier of R, the carrier of R such that A6: for x, y being set holds ( [x,y] in QOE iff ( x in the carrier of R & y in the carrier of R & S1[x,y] ) ) from RELSET_1:sch_1(); set S = RelStr(# the carrier of R,QOE #); now__::_thesis:_for_x,_y_being_set_st_[x,y]_in_the_InternalRel_of_R_holds_ [x,y]_in_QOE let x, y be set ; ::_thesis: ( [x,y] in the InternalRel of R implies [x,y] in QOE ) assume A7: [x,y] in the InternalRel of R ; ::_thesis: [x,y] in QOE A8: x in dom the InternalRel of R by A7, XTUPLE_0:def_12; y in rng the InternalRel of R by A7, XTUPLE_0:def_13; hence [x,y] in QOE by A6, A7, A8; ::_thesis: verum end; then A9: the InternalRel of R c= the InternalRel of RelStr(# the carrier of R,QOE #) by RELAT_1:def_3; A10: the InternalRel of R is_reflexive_in the carrier of R by A2, ORDERS_2:def_2; now__::_thesis:_for_x_being_set_st_x_in_the_carrier_of_R_holds_ [x,x]_in_QOE let x be set ; ::_thesis: ( x in the carrier of R implies [x,x] in QOE ) assume A11: x in the carrier of R ; ::_thesis: [x,x] in QOE [x,x] in the InternalRel of R by A10, A11, RELAT_2:def_1; hence [x,x] in QOE by A6, A11; ::_thesis: verum end; then QOE is_reflexive_in the carrier of R by RELAT_2:def_1; then A12: RelStr(# the carrier of R,QOE #) is reflexive by ORDERS_2:def_2; A13: the InternalRel of R is_transitive_in the carrier of R by A3, ORDERS_2:def_3; now__::_thesis:_for_x,_y,_z_being_set_st_x_in_the_carrier_of_R_&_y_in_the_carrier_of_R_&_z_in_the_carrier_of_R_&_[x,y]_in_QOE_&_[y,z]_in_QOE_holds_ [x,z]_in_QOE let x, y, z be set ; ::_thesis: ( x in the carrier of R & y in the carrier of R & z in the carrier of R & [x,y] in QOE & [y,z] in QOE implies [x,z] in QOE ) assume that A14: x in the carrier of R and A15: y in the carrier of R and A16: z in the carrier of R and A17: [x,y] in QOE and A18: [y,z] in QOE ; ::_thesis: [x,z] in QOE now__::_thesis:_[x,z]_in_QOE percases ( [x,y] in the InternalRel of R or ex i, j being Element of NAT st ( i < j & [x,(f . j)] in the InternalRel of R & [(f . i),y] in the InternalRel of R ) ) by A6, A17; supposeA19: [x,y] in the InternalRel of R ; ::_thesis: [x,z] in QOE now__::_thesis:_[x,z]_in_QOE percases ( [y,z] in the InternalRel of R or ex i, j being Element of NAT st ( i < j & [y,(f . j)] in the InternalRel of R & [(f . i),z] in the InternalRel of R ) ) by A6, A18; suppose [y,z] in the InternalRel of R ; ::_thesis: [x,z] in QOE then [x,z] in the InternalRel of R by A13, A14, A15, A16, A19, RELAT_2:def_8; hence [x,z] in QOE by A6, A14, A16; ::_thesis: verum end; suppose ex i, j being Element of NAT st ( i < j & [y,(f . j)] in the InternalRel of R & [(f . i),z] in the InternalRel of R ) ; ::_thesis: [x,z] in QOE then consider i, j being Element of NAT such that A20: i < j and A21: [y,(f . j)] in the InternalRel of R and A22: [(f . i),z] in the InternalRel of R ; [x,(f . j)] in the InternalRel of R by A13, A14, A15, A19, A21, RELAT_2:def_8; hence [x,z] in QOE by A6, A14, A16, A20, A22; ::_thesis: verum end; end; end; hence [x,z] in QOE ; ::_thesis: verum end; suppose ex i, j being Element of NAT st ( i < j & [x,(f . j)] in the InternalRel of R & [(f . i),y] in the InternalRel of R ) ; ::_thesis: [x,z] in QOE then consider i, j being Element of NAT such that A23: i < j and A24: [x,(f . j)] in the InternalRel of R and A25: [(f . i),y] in the InternalRel of R ; now__::_thesis:_[x,z]_in_QOE percases ( [y,z] in the InternalRel of R or ex a, b being Element of NAT st ( a < b & [y,(f . b)] in the InternalRel of R & [(f . a),z] in the InternalRel of R ) ) by A6, A18; suppose [y,z] in the InternalRel of R ; ::_thesis: [x,z] in QOE then [(f . i),z] in the InternalRel of R by A13, A15, A16, A25, RELAT_2:def_8; hence [x,z] in QOE by A6, A14, A16, A23, A24; ::_thesis: verum end; suppose ex a, b being Element of NAT st ( a < b & [y,(f . b)] in the InternalRel of R & [(f . a),z] in the InternalRel of R ) ; ::_thesis: [x,z] in QOE then consider a, b being Element of NAT such that A26: a < b and A27: [y,(f . b)] in the InternalRel of R and A28: [(f . a),z] in the InternalRel of R ; [(f . i),(f . b)] in the InternalRel of R by A13, A15, A25, A27, RELAT_2:def_8; then f . i <= f . b by ORDERS_2:def_5; then not i < b by A5; then a <= i by A26, XXREAL_0:2; then a < j by A23, XXREAL_0:2; hence [x,z] in QOE by A6, A14, A16, A24, A28; ::_thesis: verum end; end; end; hence [x,z] in QOE ; ::_thesis: verum end; end; end; hence [x,z] in QOE ; ::_thesis: verum end; then QOE is_transitive_in the carrier of R by RELAT_2:def_8; then RelStr(# the carrier of R,QOE #) is transitive by ORDERS_2:def_3; then RelStr(# the carrier of R,QOE #) is quasi_ordered by A12, Def3; then A29: RelStr(# the carrier of R,QOE #) \~ is well_founded by A4, A9; reconsider f9 = f as sequence of (RelStr(# the carrier of R,QOE #) \~) ; now__::_thesis:_for_n_being_Nat_holds_ (_f9_._(n_+_1)_<>_f9_._n_&_[(f9_._(n_+_1)),(f9_._n)]_in_the_InternalRel_of_(RelStr(#_the_carrier_of_R,QOE_#)_\~)_) let n be Nat; ::_thesis: ( f9 . (n + 1) <> f9 . n & [(f9 . (n + 1)),(f9 . n)] in the InternalRel of (RelStr(# the carrier of R,QOE #) \~) ) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; A30: n < n + 1 by NAT_1:13; then not f . n1 <= f . (n1 + 1) by A5; then A31: not [(f . n),(f . (n + 1))] in the InternalRel of R by ORDERS_2:def_5; hence f9 . (n + 1) <> f9 . n by A10, RELAT_2:def_1; ::_thesis: [(f9 . (n + 1)),(f9 . n)] in the InternalRel of (RelStr(# the carrier of R,QOE #) \~) A32: [(f9 . (n + 1)),(f9 . (n + 1))] in the InternalRel of R by A10, RELAT_2:def_1; A33: [(f9 . n),(f9 . n)] in the InternalRel of R by A10, RELAT_2:def_1; A34: now__::_thesis:_not_[(f9_._n),(f9_._(n_+_1))]_in_QOE assume [(f9 . n),(f9 . (n + 1))] in QOE ; ::_thesis: contradiction then ex i, j being Element of NAT st ( i < j & [(f9 . n),(f . j)] in the InternalRel of R & [(f . i),(f9 . (n + 1))] in the InternalRel of R ) by A6, A31; then consider l, k being Element of NAT such that A35: k < l and A36: [(f9 . n),(f . l)] in the InternalRel of R and A37: [(f . k),(f9 . (n + 1))] in the InternalRel of R ; A38: f . n <= f . l by A36, ORDERS_2:def_5; A39: f . k <= f . (n + 1) by A37, ORDERS_2:def_5; A40: l <= n1 by A5, A38; A41: n + 1 <= k by A5, A39; l < n + 1 by A40, NAT_1:13; hence contradiction by A35, A41, XXREAL_0:2; ::_thesis: verum end; A42: [(f9 . (n1 + 1)),(f9 . n1)] in QOE by A6, A30, A32, A33; not [(f9 . (n + 1)),(f9 . n)] in QOE ~ by A34, RELAT_1:def_7; hence [(f9 . (n + 1)),(f9 . n)] in the InternalRel of (RelStr(# the carrier of R,QOE #) \~) by A42, XBOOLE_0:def_5; ::_thesis: verum end; then f9 is descending by WELLFND1:def_6; hence contradiction by A29, WELLFND1:14; ::_thesis: verum end; hence R is Dickson ; ::_thesis: verum end;