:: BAGORDER semantic presentation begin theorem Th1: :: BAGORDER:1 for x, y, z being set st z in x & z in y & x \ {z} = y \ {z} holds x = y proof let x, y, z be set ; ::_thesis: ( z in x & z in y & x \ {z} = y \ {z} implies x = y ) assume that A1: z in x and A2: z in y ; ::_thesis: ( not x \ {z} = y \ {z} or x = y ) assume A3: x \ {z} = y \ {z} ; ::_thesis: x = y thus x = x \/ {z} by A1, ZFMISC_1:40 .= (y \ {z}) \/ {z} by A3, XBOOLE_1:39 .= y \/ {z} by XBOOLE_1:39 .= y by A2, ZFMISC_1:40 ; ::_thesis: verum end; theorem :: BAGORDER:2 for n, k being Element of NAT holds ( k in Seg n iff ( k - 1 is Element of NAT & k - 1 < n ) ) proof let n, k be Element of NAT ; ::_thesis: ( k in Seg n iff ( k - 1 is Element of NAT & k - 1 < n ) ) A1: Seg n = { x where x is Element of NAT : ( 1 <= x & x <= n ) } by FINSEQ_1:def_1; hereby ::_thesis: ( k - 1 is Element of NAT & k - 1 < n implies k in Seg n ) assume k in Seg n ; ::_thesis: ( k - 1 is Element of NAT & k - 1 < n ) then consider x being Element of NAT such that A2: k = x and A3: 1 <= x and A4: x <= n by A1; set x1 = k - 1; set n1 = n - 1; 0 < x by A3; then reconsider x1 = k - 1 as Element of NAT by A2, NAT_1:20; x1 = k - 1 ; hence k - 1 is Element of NAT ; ::_thesis: k - 1 < n 0 < n by A3, A4; then reconsider n1 = n - 1 as Element of NAT by NAT_1:20; k + (- 1) <= n + (- 1) by A2, A4, XREAL_1:6; then x1 <= n1 ; then k - 1 < n1 + 1 by NAT_1:13; hence k - 1 < n ; ::_thesis: verum end; assume that A5: k - 1 is Element of NAT and A6: k - 1 < n ; ::_thesis: k in Seg n reconsider k1 = k - 1 as Element of NAT by A5; 0 <= k1 ; then A7: 0 + 1 <= (k - 1) + 1 by XREAL_1:6; (k - 1) + 1 <= (n - 1) + 1 by A5, A6, NAT_1:13; hence k in Seg n by A1, A7; ::_thesis: verum end; registration let f be finite-support Function; let X be set ; clusterf | X -> finite-support ; coherence f | X is finite-support proof support (f | X) c= support f proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in support (f | X) or x in support f ) assume A1: x in support (f | X) ; ::_thesis: x in support f then A2: (f | X) . x <> 0 by PRE_POLY:def_7; support (f | X) c= dom (f | X) by PRE_POLY:37; then f . x <> 0 by A1, A2, FUNCT_1:47; hence x in support f by PRE_POLY:def_7; ::_thesis: verum end; hence f | X is finite-support by PRE_POLY:def_8; ::_thesis: verum end; end; theorem :: BAGORDER:3 canceled; theorem Th4: :: BAGORDER:4 for fs being FinSequence of NAT holds ( Sum fs = 0 iff fs = (len fs) |-> 0 ) proof let fs be FinSequence of NAT ; ::_thesis: ( Sum fs = 0 iff fs = (len fs) |-> 0 ) hereby ::_thesis: ( fs = (len fs) |-> 0 implies Sum fs = 0 ) assume A1: Sum fs = 0 ; ::_thesis: fs = (len fs) |-> 0 A2: Seg (len fs) = dom fs by FINSEQ_1:def_3; A3: Seg (len fs) = dom ((len fs) |-> 0) by FUNCOP_1:13; now__::_thesis:_for_k_being_Nat_st_k_in_Seg_(len_fs)_holds_ fs_._k_=_((len_fs)_|->_0)_._k let k be Nat; ::_thesis: ( k in Seg (len fs) implies fs . k = ((len fs) |-> 0) . k ) assume A4: k in Seg (len fs) ; ::_thesis: fs . k = ((len fs) |-> 0) . k now__::_thesis:_not_fs_._k_<>_0 assume A5: fs . k <> 0 ; ::_thesis: contradiction for k being Nat st k in dom fs holds 0 <= fs . k ; hence contradiction by A1, A2, A4, A5, RVSUM_1:85; ::_thesis: verum end; hence fs . k = ((len fs) |-> 0) . k by A4, FUNCOP_1:7; ::_thesis: verum end; hence fs = (len fs) |-> 0 by A2, A3, FINSEQ_1:13; ::_thesis: verum end; assume fs = (len fs) |-> 0 ; ::_thesis: Sum fs = 0 hence Sum fs = 0 by RVSUM_1:81; ::_thesis: verum end; definition let n, i, j be Nat; let b be ManySortedSet of n; func(i,j) -cut b -> ManySortedSet of j -' i means :Def1: :: BAGORDER:def 1 for k being Element of NAT st k in j -' i holds it . k = b . (i + k); existence ex b1 being ManySortedSet of j -' i st for k being Element of NAT st k in j -' i holds b1 . k = b . (i + k) proof defpred S1[ set , set ] means ex k1 being Element of NAT st ( k1 = $1 & $2 = b . (i + k1) ); now__::_thesis:_for_x_being_set_st_x_in_j_-'_i_holds_ ex_y_being_set_st_S1[x,y] let x be set ; ::_thesis: ( x in j -' i implies ex y being set st S1[x,y] ) assume A1: x in j -' i ; ::_thesis: ex y being set st S1[x,y] j -' i = { k where k is Element of NAT : k < j -' i } by AXIOMS:4; then ex k being Element of NAT st ( x = k & k < j -' i ) by A1; then reconsider x9 = x as Element of NAT ; consider y being set such that A2: y = b . (i + x9) ; take y = y; ::_thesis: S1[x,y] thus S1[x,y] by A2; ::_thesis: verum end; then A3: for x being set st x in j -' i holds ex y being set st S1[x,y] ; consider f being Function such that A4: dom f = j -' i and A5: for k being set st k in j -' i holds S1[k,f . k] from CLASSES1:sch_1(A3); reconsider f = f as ManySortedSet of j -' i by A4, PARTFUN1:def_2, RELAT_1:def_18; take f ; ::_thesis: for k being Element of NAT st k in j -' i holds f . k = b . (i + k) let k be Element of NAT ; ::_thesis: ( k in j -' i implies f . k = b . (i + k) ) assume k in j -' i ; ::_thesis: f . k = b . (i + k) then ex k9 being Element of NAT st ( k9 = k & f . k = b . (i + k9) ) by A5; hence f . k = b . (i + k) ; ::_thesis: verum end; uniqueness for b1, b2 being ManySortedSet of j -' i st ( for k being Element of NAT st k in j -' i holds b1 . k = b . (i + k) ) & ( for k being Element of NAT st k in j -' i holds b2 . k = b . (i + k) ) holds b1 = b2 proof let IT1, IT2 be ManySortedSet of j -' i; ::_thesis: ( ( for k being Element of NAT st k in j -' i holds IT1 . k = b . (i + k) ) & ( for k being Element of NAT st k in j -' i holds IT2 . k = b . (i + k) ) implies IT1 = IT2 ) assume that A6: for k being Element of NAT st k in j -' i holds IT1 . k = b . (i + k) and A7: for k being Element of NAT st k in j -' i holds IT2 . k = b . (i + k) ; ::_thesis: IT1 = IT2 A8: j -' i = dom IT1 by PARTFUN1:def_2; A9: j -' i = dom IT2 by PARTFUN1:def_2; now__::_thesis:_for_x_being_set_st_x_in_j_-'_i_holds_ IT1_._x_=_IT2_._x let x be set ; ::_thesis: ( x in j -' i implies IT1 . x = IT2 . x ) assume A10: x in j -' i ; ::_thesis: IT1 . x = IT2 . x j -' i = { k where k is Element of NAT : k < j -' i } by AXIOMS:4; then ex k being Element of NAT st ( x = k & k < j -' i ) by A10; then reconsider x9 = x as Element of NAT ; IT1 . x = b . (i + x9) by A6, A10; hence IT1 . x = IT2 . x by A7, A10; ::_thesis: verum end; hence IT1 = IT2 by A8, A9, FUNCT_1:2; ::_thesis: verum end; end; :: deftheorem Def1 defines -cut BAGORDER:def_1_:_ for n, i, j being Nat for b being ManySortedSet of n for b5 being ManySortedSet of j -' i holds ( b5 = (i,j) -cut b iff for k being Element of NAT st k in j -' i holds b5 . k = b . (i + k) ); registration let n, i, j be Nat; let b be natural-valued ManySortedSet of n; cluster(i,j) -cut b -> natural-valued ; coherence (i,j) -cut b is natural-valued proof now__::_thesis:_for_y_being_set_st_y_in_rng_((i,j)_-cut_b)_holds_ y_in_NAT let y be set ; ::_thesis: ( y in rng ((i,j) -cut b) implies y in NAT ) assume y in rng ((i,j) -cut b) ; ::_thesis: y in NAT then consider x being set such that A1: x in dom ((i,j) -cut b) and A2: ((i,j) -cut b) . x = y by FUNCT_1:def_3; A3: x in j -' i by A1; j -' i = { k where k is Element of NAT : k < j -' i } by AXIOMS:4; then ex k being Element of NAT st ( k = x & k < j -' i ) by A3; then reconsider x = x as Element of NAT ; y = b . (i + x) by A2, A3, Def1; hence y in NAT ; ::_thesis: verum end; then rng ((i,j) -cut b) c= NAT by TARSKI:def_3; hence (i,j) -cut b is natural-valued by VALUED_0:def_6; ::_thesis: verum end; end; registration let n, i, j be Element of NAT ; let b be finite-support ManySortedSet of n; cluster(i,j) -cut b -> finite-support ; coherence (i,j) -cut b is finite-support ; end; theorem Th5: :: BAGORDER:5 for n, i being Nat for a, b being ManySortedSet of n holds ( a = b iff ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & ((i + 1),n) -cut a = ((i + 1),n) -cut b ) ) proof let n, i be Nat; ::_thesis: for a, b being ManySortedSet of n holds ( a = b iff ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & ((i + 1),n) -cut a = ((i + 1),n) -cut b ) ) let a, b be ManySortedSet of n; ::_thesis: ( a = b iff ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & ((i + 1),n) -cut a = ((i + 1),n) -cut b ) ) set CUTA1 = (0,(i + 1)) -cut a; set CUTA2 = ((i + 1),n) -cut a; set CUTB1 = (0,(i + 1)) -cut b; set CUTB2 = ((i + 1),n) -cut b; thus ( a = b implies ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & ((i + 1),n) -cut a = ((i + 1),n) -cut b ) ) ; ::_thesis: ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & ((i + 1),n) -cut a = ((i + 1),n) -cut b implies a = b ) assume that A1: (0,(i + 1)) -cut a = (0,(i + 1)) -cut b and A2: ((i + 1),n) -cut a = ((i + 1),n) -cut b ; ::_thesis: a = b A3: now__::_thesis:_for_k_being_Element_of_NAT_st_k_in_i_+_1_holds_ a_._k_=_b_._k let k be Element of NAT ; ::_thesis: ( k in i + 1 implies a . k = b . k ) assume A4: k in i + 1 ; ::_thesis: a . k = b . k (i + 1) -' 0 = ((i + 1) + 0) -' 0 ; then A5: k in (i + 1) -' 0 by A4, NAT_D:34; then ((0,(i + 1)) -cut b) . k = b . (0 + k) by Def1; hence a . k = b . k by A1, A5, Def1; ::_thesis: verum end; A6: now__::_thesis:_for_x_being_Element_of_NAT_st_x_>=_i_+_1_&_x_<_n_holds_ a_._x_=_b_._x let x be Element of NAT ; ::_thesis: ( x >= i + 1 & x < n implies a . x = b . x ) assume that A7: x >= i + 1 and A8: x < n ; ::_thesis: a . x = b . x set k = x -' (i + 1); x - (i + 1) >= (i + 1) - (i + 1) by A7, XREAL_1:9; then A9: x -' (i + 1) = x - (i + 1) by XREAL_0:def_2; n >= i + 1 by A7, A8, XXREAL_0:2; then n - (i + 1) >= (i + 1) - (i + 1) by XREAL_1:9; then A10: n -' (i + 1) = n - (i + 1) by XREAL_0:def_2; x - (i + 1) < n - (i + 1) by A8, XREAL_1:14; then A11: x -' (i + 1) in n -' (i + 1) by A9, A10, NAT_1:44; then (((i + 1),n) -cut b) . (x -' (i + 1)) = b . ((i + 1) + (x -' (i + 1))) by Def1; hence a . x = b . x by A2, A9, A11, Def1; ::_thesis: verum end; now__::_thesis:_for_x9_being_set_st_x9_in_n_holds_ a_._x9_=_b_._x9 let x9 be set ; ::_thesis: ( x9 in n implies a . b1 = b . b1 ) assume A12: x9 in n ; ::_thesis: a . b1 = b . b1 n = { k where k is Element of NAT : k < n } by AXIOMS:4; then A13: ex k being Element of NAT st ( k = x9 & k < n ) by A12; then reconsider x = x9 as Element of NAT ; percases ( x in i + 1 or not x in i + 1 ) ; suppose x in i + 1 ; ::_thesis: a . b1 = b . b1 hence a . x9 = b . x9 by A3; ::_thesis: verum end; suppose not x in i + 1 ; ::_thesis: a . b1 = b . b1 then x >= i + 1 by NAT_1:44; hence a . x9 = b . x9 by A6, A13; ::_thesis: verum end; end; end; hence a = b by PBOOLE:3; ::_thesis: verum end; definition let x be non empty set ; let n be non empty Element of NAT ; func Fin (x,n) -> set equals :: BAGORDER:def 2 { y where y is Subset of x : ( y is finite & not y is empty & card y c= n ) } ; coherence { y where y is Subset of x : ( y is finite & not y is empty & card y c= n ) } is set ; end; :: deftheorem defines Fin BAGORDER:def_2_:_ for x being non empty set for n being non empty Element of NAT holds Fin (x,n) = { y where y is Subset of x : ( y is finite & not y is empty & card y c= n ) } ; registration let x be non empty set ; let n be non empty Element of NAT ; cluster Fin (x,n) -> non empty ; coherence not Fin (x,n) is empty proof set y = the Element of x; A1: 0 + 1 < n + 1 by XREAL_1:8; now__::_thesis:_card_{_the_Element_of_x}_c=_n percases ( 1 c= n or n in 1 ) by ORDINAL1:16; suppose 1 c= n ; ::_thesis: card { the Element of x} c= n hence card { the Element of x} c= n by CARD_1:30; ::_thesis: verum end; supposeA2: n in 1 ; ::_thesis: card { the Element of x} c= n 1 = { k where k is Element of NAT : k < 1 } by AXIOMS:4; then ex k being Element of NAT st ( k = n & k < 1 ) by A2; hence card { the Element of x} c= n by A1, NAT_1:13; ::_thesis: verum end; end; end; then { the Element of x} in Fin (x,n) ; hence not Fin (x,n) is empty ; ::_thesis: verum end; end; theorem Th6: :: BAGORDER:6 for R being non empty transitive antisymmetric RelStr for X being finite Subset of R st X <> {} holds ex x being Element of R st ( x in X & x is_maximal_wrt X, the InternalRel of R ) proof let R be non empty transitive antisymmetric RelStr ; ::_thesis: for X being finite Subset of R st X <> {} holds ex x being Element of R st ( x in X & x is_maximal_wrt X, the InternalRel of R ) let X be finite Subset of R; ::_thesis: ( X <> {} implies ex x being Element of R st ( x in X & x is_maximal_wrt X, the InternalRel of R ) ) set IR = the InternalRel of R; set CR = the carrier of R; A1: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; A2: the InternalRel of R is_antisymmetric_in the carrier of R by ORDERS_2:def_4; A3: X is finite ; defpred S1[ set ] means ( $1 <> {} implies ex x being Element of R st ( x in $1 & x is_maximal_wrt $1, the InternalRel of R ) ); A4: S1[ {} ] ; now__::_thesis:_for_y,_B_being_set_st_y_in_X_&_B_c=_X_&_(_B_<>_{}_implies_ex_x_being_Element_of_R_st_ (_x_in_B_&_x_is_maximal_wrt_B,_the_InternalRel_of_R_)_)_&_B_\/_{y}_<>_{}_holds_ ex_y9_being_Element_of_R_st_ (_y9_in_B_\/_{y}_&_y9_is_maximal_wrt_B_\/_{y},_the_InternalRel_of_R_) let y, B be set ; ::_thesis: ( y in X & B c= X & ( B <> {} implies ex x being Element of R st ( x in B & x is_maximal_wrt B, the InternalRel of R ) ) & B \/ {y} <> {} implies ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_maximal_wrt b2 \/ {y9}, the InternalRel of R ) ) assume that A5: y in X and B c= X and A6: ( B <> {} implies ex x being Element of R st ( x in B & x is_maximal_wrt B, the InternalRel of R ) ) ; ::_thesis: ( B \/ {y} <> {} implies ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_maximal_wrt b2 \/ {y9}, the InternalRel of R ) ) reconsider y9 = y as Element of R by A5; assume B \/ {y} <> {} ; ::_thesis: ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_maximal_wrt b2 \/ {y9}, the InternalRel of R ) percases ( B = {} or B <> {} ) ; supposeA7: B = {} ; ::_thesis: ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_maximal_wrt b2 \/ {y9}, the InternalRel of R ) take y9 = y9; ::_thesis: ( y9 in B \/ {y} & y9 is_maximal_wrt B \/ {y}, the InternalRel of R ) thus y9 in B \/ {y} by A7, TARSKI:def_1; ::_thesis: y9 is_maximal_wrt B \/ {y}, the InternalRel of R A8: y9 in B \/ {y} by A7, TARSKI:def_1; for z being set holds ( not z in B \/ {y9} or not z <> y9 or not [y9,z] in the InternalRel of R ) by A7, TARSKI:def_1; hence y9 is_maximal_wrt B \/ {y}, the InternalRel of R by A8, WAYBEL_4:def_23; ::_thesis: verum end; suppose B <> {} ; ::_thesis: ex x being Element of R st ( b3 in b2 \/ {x} & b3 is_maximal_wrt b2 \/ {x}, the InternalRel of R ) then consider x being Element of R such that A9: x in B and A10: x is_maximal_wrt B, the InternalRel of R by A6; now__::_thesis:_ex_y9_being_Element_of_R_st_ (_y9_in_B_\/_{y}_&_y9_is_maximal_wrt_B_\/_{y},_the_InternalRel_of_R_) percases ( [x,y] in the InternalRel of R or [y,x] in the InternalRel of R or ( not [x,y] in the InternalRel of R & not [y,x] in the InternalRel of R ) ) ; supposeA11: [x,y] in the InternalRel of R ; ::_thesis: ex y9 being Element of R st ( y9 in B \/ {y} & y9 is_maximal_wrt B \/ {y}, the InternalRel of R ) take y9 = y9; ::_thesis: ( y9 in B \/ {y} & y9 is_maximal_wrt B \/ {y}, the InternalRel of R ) A12: y in {y} by TARSKI:def_1; hence y9 in B \/ {y} by XBOOLE_0:def_3; ::_thesis: y9 is_maximal_wrt B \/ {y}, the InternalRel of R A13: now__::_thesis:_for_z_being_set_holds_ (_not_z_in_B_\/_{y}_or_not_z_<>_y_or_not_[y,z]_in_the_InternalRel_of_R_) given z being set such that A14: z in B \/ {y} and A15: z <> y and A16: [y,z] in the InternalRel of R ; ::_thesis: contradiction A17: y9 in the carrier of R ; z in the carrier of R by A16, ZFMISC_1:87; then A18: [x,z] in the InternalRel of R by A1, A11, A16, A17, RELAT_2:def_8; percases ( z in B or z in {y} ) by A14, XBOOLE_0:def_3; supposeA19: z in B ; ::_thesis: contradiction now__::_thesis:_contradiction percases ( z = x or z <> x ) ; supposeA20: z = x ; ::_thesis: contradiction then x = y9 by A2, A11, A16, RELAT_2:def_4; hence contradiction by A15, A20; ::_thesis: verum end; suppose z <> x ; ::_thesis: contradiction hence contradiction by A10, A18, A19, WAYBEL_4:def_23; ::_thesis: verum end; end; end; hence contradiction ; ::_thesis: verum end; suppose z in {y} ; ::_thesis: contradiction hence contradiction by A15, TARSKI:def_1; ::_thesis: verum end; end; end; y9 in B \/ {y} by A12, XBOOLE_0:def_3; hence y9 is_maximal_wrt B \/ {y}, the InternalRel of R by A13, WAYBEL_4:def_23; ::_thesis: verum end; supposeA21: [y,x] in the InternalRel of R ; ::_thesis: ex x being Element of R st ( x in B \/ {y} & x is_maximal_wrt B \/ {y}, the InternalRel of R ) take x = x; ::_thesis: ( x in B \/ {y} & x is_maximal_wrt B \/ {y}, the InternalRel of R ) thus x in B \/ {y} by A9, XBOOLE_0:def_3; ::_thesis: x is_maximal_wrt B \/ {y}, the InternalRel of R A22: now__::_thesis:_for_z_being_set_holds_ (_not_z_in_B_\/_{y}_or_not_z_<>_x_or_not_[x,z]_in_the_InternalRel_of_R_) assume ex z being set st ( z in B \/ {y} & z <> x & [x,z] in the InternalRel of R ) ; ::_thesis: contradiction then consider z being set such that A23: z in B \/ {y} and A24: z <> x and A25: [x,z] in the InternalRel of R ; percases ( z in B or z in {y} ) by A23, XBOOLE_0:def_3; suppose z in B ; ::_thesis: contradiction hence contradiction by A10, A24, A25, WAYBEL_4:def_23; ::_thesis: verum end; suppose z in {y} ; ::_thesis: contradiction then A26: z = y by TARSKI:def_1; z in the carrier of R by A25, ZFMISC_1:87; hence contradiction by A2, A21, A24, A25, A26, RELAT_2:def_4; ::_thesis: verum end; end; end; x in B \/ {y} by A9, XBOOLE_0:def_3; hence x is_maximal_wrt B \/ {y}, the InternalRel of R by A22, WAYBEL_4:def_23; ::_thesis: verum end; supposeA27: ( not [x,y] in the InternalRel of R & not [y,x] in the InternalRel of R ) ; ::_thesis: ex x being Element of R st ( x in B \/ {y} & x is_maximal_wrt B \/ {y}, the InternalRel of R ) take x = x; ::_thesis: ( x in B \/ {y} & x is_maximal_wrt B \/ {y}, the InternalRel of R ) thus x in B \/ {y} by A9, XBOOLE_0:def_3; ::_thesis: x is_maximal_wrt B \/ {y}, the InternalRel of R A28: now__::_thesis:_for_z_being_set_holds_ (_not_z_in_B_\/_{y}_or_not_z_<>_x_or_not_[x,z]_in_the_InternalRel_of_R_) assume ex z being set st ( z in B \/ {y} & z <> x & [x,z] in the InternalRel of R ) ; ::_thesis: contradiction then consider z being set such that A29: z in B \/ {y} and A30: z <> x and A31: [x,z] in the InternalRel of R ; percases ( z in B or z in {y} ) by A29, XBOOLE_0:def_3; suppose z in B ; ::_thesis: contradiction hence contradiction by A10, A30, A31, WAYBEL_4:def_23; ::_thesis: verum end; suppose z in {y} ; ::_thesis: contradiction hence contradiction by A27, A31, TARSKI:def_1; ::_thesis: verum end; end; end; x in B \/ {y} by A9, XBOOLE_0:def_3; hence x is_maximal_wrt B \/ {y}, the InternalRel of R by A28, WAYBEL_4:def_23; ::_thesis: verum end; end; end; hence ex x being Element of R st ( x in B \/ {y} & x is_maximal_wrt B \/ {y}, the InternalRel of R ) ; ::_thesis: verum end; end; end; then A32: for y, B being set st y in X & B c= X & S1[B] holds S1[B \/ {y}] ; thus S1[X] from FINSET_1:sch_2(A3, A4, A32); ::_thesis: verum end; theorem Th7: :: BAGORDER:7 for R being non empty transitive antisymmetric RelStr for X being finite Subset of R st X <> {} holds ex x being Element of R st ( x in X & x is_minimal_wrt X, the InternalRel of R ) proof let R be non empty transitive antisymmetric RelStr ; ::_thesis: for X being finite Subset of R st X <> {} holds ex x being Element of R st ( x in X & x is_minimal_wrt X, the InternalRel of R ) let X be finite Subset of R; ::_thesis: ( X <> {} implies ex x being Element of R st ( x in X & x is_minimal_wrt X, the InternalRel of R ) ) set IR = the InternalRel of R; set CR = the carrier of R; A1: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; A2: the InternalRel of R is_antisymmetric_in the carrier of R by ORDERS_2:def_4; A3: X is finite ; defpred S1[ set ] means ( $1 <> {} implies ex x being Element of R st ( x in $1 & x is_minimal_wrt $1, the InternalRel of R ) ); A4: S1[ {} ] ; now__::_thesis:_for_y,_B_being_set_st_y_in_X_&_B_c=_X_&_(_B_<>_{}_implies_ex_x_being_Element_of_R_st_ (_x_in_B_&_x_is_minimal_wrt_B,_the_InternalRel_of_R_)_)_&_B_\/_{y}_<>_{}_holds_ ex_y9_being_Element_of_R_st_ (_y9_in_B_\/_{y}_&_y9_is_minimal_wrt_B_\/_{y},_the_InternalRel_of_R_) let y, B be set ; ::_thesis: ( y in X & B c= X & ( B <> {} implies ex x being Element of R st ( x in B & x is_minimal_wrt B, the InternalRel of R ) ) & B \/ {y} <> {} implies ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_minimal_wrt b2 \/ {y9}, the InternalRel of R ) ) assume that A5: y in X and B c= X and A6: ( B <> {} implies ex x being Element of R st ( x in B & x is_minimal_wrt B, the InternalRel of R ) ) ; ::_thesis: ( B \/ {y} <> {} implies ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_minimal_wrt b2 \/ {y9}, the InternalRel of R ) ) reconsider y9 = y as Element of R by A5; assume B \/ {y} <> {} ; ::_thesis: ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_minimal_wrt b2 \/ {y9}, the InternalRel of R ) percases ( B = {} or B <> {} ) ; supposeA7: B = {} ; ::_thesis: ex y9 being Element of R st ( b3 in b2 \/ {y9} & b3 is_minimal_wrt b2 \/ {y9}, the InternalRel of R ) take y9 = y9; ::_thesis: ( y9 in B \/ {y} & y9 is_minimal_wrt B \/ {y}, the InternalRel of R ) thus y9 in B \/ {y} by A7, TARSKI:def_1; ::_thesis: y9 is_minimal_wrt B \/ {y}, the InternalRel of R A8: y9 in B \/ {y} by A7, TARSKI:def_1; for z being set holds ( not z in B \/ {y9} or not z <> y9 or not [z,y9] in the InternalRel of R ) by A7, TARSKI:def_1; hence y9 is_minimal_wrt B \/ {y}, the InternalRel of R by A8, WAYBEL_4:def_25; ::_thesis: verum end; suppose B <> {} ; ::_thesis: ex x being Element of R st ( b3 in b2 \/ {x} & b3 is_minimal_wrt b2 \/ {x}, the InternalRel of R ) then consider x being Element of R such that A9: x in B and A10: x is_minimal_wrt B, the InternalRel of R by A6; now__::_thesis:_ex_y9_being_Element_of_R_st_ (_y9_in_B_\/_{y}_&_y9_is_minimal_wrt_B_\/_{y},_the_InternalRel_of_R_) percases ( [y,x] in the InternalRel of R or [x,y] in the InternalRel of R or ( not [x,y] in the InternalRel of R & not [y,x] in the InternalRel of R ) ) ; supposeA11: [y,x] in the InternalRel of R ; ::_thesis: ex y9 being Element of R st ( y9 in B \/ {y} & y9 is_minimal_wrt B \/ {y}, the InternalRel of R ) take y9 = y9; ::_thesis: ( y9 in B \/ {y} & y9 is_minimal_wrt B \/ {y}, the InternalRel of R ) A12: y in {y} by TARSKI:def_1; hence y9 in B \/ {y} by XBOOLE_0:def_3; ::_thesis: y9 is_minimal_wrt B \/ {y}, the InternalRel of R A13: now__::_thesis:_for_z_being_set_holds_ (_not_z_in_B_\/_{y}_or_not_z_<>_y_or_not_[z,y]_in_the_InternalRel_of_R_) assume ex z being set st ( z in B \/ {y} & z <> y & [z,y] in the InternalRel of R ) ; ::_thesis: contradiction then consider z being set such that A14: z in B \/ {y} and A15: z <> y and A16: [z,y] in the InternalRel of R ; A17: y9 in the carrier of R ; z in the carrier of R by A16, ZFMISC_1:87; then A18: [z,x] in the InternalRel of R by A1, A11, A16, A17, RELAT_2:def_8; percases ( z in B or z in {y} ) by A14, XBOOLE_0:def_3; supposeA19: z in B ; ::_thesis: contradiction now__::_thesis:_contradiction percases ( z = x or z <> x ) ; supposeA20: z = x ; ::_thesis: contradiction then x = y9 by A2, A11, A16, RELAT_2:def_4; hence contradiction by A15, A20; ::_thesis: verum end; suppose z <> x ; ::_thesis: contradiction hence contradiction by A10, A18, A19, WAYBEL_4:def_25; ::_thesis: verum end; end; end; hence contradiction ; ::_thesis: verum end; suppose z in {y} ; ::_thesis: contradiction hence contradiction by A15, TARSKI:def_1; ::_thesis: verum end; end; end; y9 in B \/ {y} by A12, XBOOLE_0:def_3; hence y9 is_minimal_wrt B \/ {y}, the InternalRel of R by A13, WAYBEL_4:def_25; ::_thesis: verum end; supposeA21: [x,y] in the InternalRel of R ; ::_thesis: ex x being Element of R st ( x in B \/ {y} & x is_minimal_wrt B \/ {y}, the InternalRel of R ) take x = x; ::_thesis: ( x in B \/ {y} & x is_minimal_wrt B \/ {y}, the InternalRel of R ) thus x in B \/ {y} by A9, XBOOLE_0:def_3; ::_thesis: x is_minimal_wrt B \/ {y}, the InternalRel of R A22: now__::_thesis:_for_z_being_set_holds_ (_not_z_in_B_\/_{y}_or_not_z_<>_x_or_not_[z,x]_in_the_InternalRel_of_R_) assume ex z being set st ( z in B \/ {y} & z <> x & [z,x] in the InternalRel of R ) ; ::_thesis: contradiction then consider z being set such that A23: z in B \/ {y} and A24: z <> x and A25: [z,x] in the InternalRel of R ; percases ( z in B or z in {y} ) by A23, XBOOLE_0:def_3; suppose z in B ; ::_thesis: contradiction hence contradiction by A10, A24, A25, WAYBEL_4:def_25; ::_thesis: verum end; suppose z in {y} ; ::_thesis: contradiction then A26: z = y by TARSKI:def_1; z in the carrier of R by A25, ZFMISC_1:87; hence contradiction by A2, A21, A24, A25, A26, RELAT_2:def_4; ::_thesis: verum end; end; end; x in B \/ {y} by A9, XBOOLE_0:def_3; hence x is_minimal_wrt B \/ {y}, the InternalRel of R by A22, WAYBEL_4:def_25; ::_thesis: verum end; supposeA27: ( not [x,y] in the InternalRel of R & not [y,x] in the InternalRel of R ) ; ::_thesis: ex x being Element of R st ( x in B \/ {y} & x is_minimal_wrt B \/ {y}, the InternalRel of R ) take x = x; ::_thesis: ( x in B \/ {y} & x is_minimal_wrt B \/ {y}, the InternalRel of R ) thus x in B \/ {y} by A9, XBOOLE_0:def_3; ::_thesis: x is_minimal_wrt B \/ {y}, the InternalRel of R A28: now__::_thesis:_for_z_being_set_holds_ (_not_z_in_B_\/_{y}_or_not_z_<>_x_or_not_[z,x]_in_the_InternalRel_of_R_) assume ex z being set st ( z in B \/ {y} & z <> x & [z,x] in the InternalRel of R ) ; ::_thesis: contradiction then consider z being set such that A29: z in B \/ {y} and A30: z <> x and A31: [z,x] in the InternalRel of R ; percases ( z in B or z in {y} ) by A29, XBOOLE_0:def_3; suppose z in B ; ::_thesis: contradiction hence contradiction by A10, A30, A31, WAYBEL_4:def_25; ::_thesis: verum end; suppose z in {y} ; ::_thesis: contradiction hence contradiction by A27, A31, TARSKI:def_1; ::_thesis: verum end; end; end; x in B \/ {y} by A9, XBOOLE_0:def_3; hence x is_minimal_wrt B \/ {y}, the InternalRel of R by A28, WAYBEL_4:def_25; ::_thesis: verum end; end; end; hence ex x being Element of R st ( x in B \/ {y} & x is_minimal_wrt B \/ {y}, the InternalRel of R ) ; ::_thesis: verum end; end; end; then A32: for y, B being set st y in X & B c= X & S1[B] holds S1[B \/ {y}] ; thus S1[X] from FINSET_1:sch_2(A3, A4, A32); ::_thesis: verum end; theorem :: BAGORDER:8 for R being non empty transitive antisymmetric RelStr for f being sequence of R st f is descending holds for j, i being Nat st i < j holds ( f . i <> f . j & [(f . j),(f . i)] in the InternalRel of R ) proof let R be non empty transitive antisymmetric RelStr ; ::_thesis: for f being sequence of R st f is descending holds for j, i being Nat st i < j holds ( f . i <> f . j & [(f . j),(f . i)] in the InternalRel of R ) let f be sequence of R; ::_thesis: ( f is descending implies for j, i being Nat st i < j holds ( f . i <> f . j & [(f . j),(f . i)] in the InternalRel of R ) ) assume A1: f is descending ; ::_thesis: for j, i being Nat st i < j holds ( f . i <> f . j & [(f . j),(f . i)] in the InternalRel of R ) set IR = the InternalRel of R; set CR = the carrier of R; A2: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; A3: the InternalRel of R is_antisymmetric_in the carrier of R by ORDERS_2:def_4; defpred S1[ Nat] means for i being Nat st i < $1 holds ( f . i <> f . $1 & [(f . $1),(f . i)] in the InternalRel of R ); A4: S1[ 0 ] ; now__::_thesis:_for_j_being_Nat_st_(_for_i_being_Nat_st_i_<_j_holds_ (_f_._i_<>_f_._j_&_[(f_._j),(f_._i)]_in_the_InternalRel_of_R_)_)_holds_ for_i_being_Nat_st_i_<_j_+_1_holds_ (_f_._i_<>_f_._(j_+_1)_&_[(f_._(j_+_1)),(f_._i)]_in_the_InternalRel_of_R_) let j be Nat; ::_thesis: ( ( for i being Nat st i < j holds ( f . i <> f . j & [(f . j),(f . i)] in the InternalRel of R ) ) implies for i being Nat st i < j + 1 holds ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) ) assume A5: for i being Nat st i < j holds ( f . i <> f . j & [(f . j),(f . i)] in the InternalRel of R ) ; ::_thesis: for i being Nat st i < j + 1 holds ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) let i be Nat; ::_thesis: ( i < j + 1 implies ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) ) assume A6: i < j + 1 ; ::_thesis: ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) now__::_thesis:_(_f_._i_<>_f_._(j_+_1)_&_[(f_._(j_+_1)),(f_._i)]_in_the_InternalRel_of_R_) percases ( i > j or i = j or i < j ) by XXREAL_0:1; suppose i > j ; ::_thesis: ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) hence ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) by A6, NAT_1:13; ::_thesis: verum end; suppose i = j ; ::_thesis: ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) hence ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) by A1, WELLFND1:def_6; ::_thesis: verum end; suppose i < j ; ::_thesis: ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) then A7: [(f . j),(f . i)] in the InternalRel of R by A5; A8: f . (j + 1) <> f . j by A1, WELLFND1:def_6; [(f . (j + 1)),(f . j)] in the InternalRel of R by A1, WELLFND1:def_6; hence ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) by A2, A3, A7, A8, RELAT_2:def_4, RELAT_2:def_8; ::_thesis: verum end; end; end; hence ( f . i <> f . (j + 1) & [(f . (j + 1)),(f . i)] in the InternalRel of R ) ; ::_thesis: verum end; then A9: for j being Nat st S1[j] holds S1[j + 1] ; thus for j being Nat holds S1[j] from NAT_1:sch_2(A4, A9); ::_thesis: verum end; definition let R be non empty RelStr ; let s be sequence of R; attrs is non-increasing means :Def3: :: BAGORDER:def 3 for i being Nat holds [(s . (i + 1)),(s . i)] in the InternalRel of R; end; :: deftheorem Def3 defines non-increasing BAGORDER:def_3_:_ for R being non empty RelStr for s being sequence of R holds ( s is non-increasing iff for i being Nat holds [(s . (i + 1)),(s . i)] in the InternalRel of R ); theorem Th9: :: BAGORDER:9 for R being non empty transitive RelStr for f being sequence of R st f is non-increasing holds for j, i being Nat st i < j holds [(f . j),(f . i)] in the InternalRel of R proof let R be non empty transitive RelStr ; ::_thesis: for f being sequence of R st f is non-increasing holds for j, i being Nat st i < j holds [(f . j),(f . i)] in the InternalRel of R let f be sequence of R; ::_thesis: ( f is non-increasing implies for j, i being Nat st i < j holds [(f . j),(f . i)] in the InternalRel of R ) assume A1: f is non-increasing ; ::_thesis: for j, i being Nat st i < j holds [(f . j),(f . i)] in the InternalRel of R set IR = the InternalRel of R; set CR = the carrier of R; A2: the InternalRel of R is_transitive_in the carrier of R by ORDERS_2:def_3; defpred S1[ Nat] means for i being Nat st i < $1 holds [(f . $1),(f . i)] in the InternalRel of R; A3: S1[ 0 ] ; now__::_thesis:_for_j_being_Nat_st_(_for_i_being_Nat_st_i_<_j_holds_ [(f_._j),(f_._i)]_in_the_InternalRel_of_R_)_holds_ for_i_being_Nat_st_i_<_j_+_1_holds_ [(f_._(j_+_1)),(f_._i)]_in_the_InternalRel_of_R let j be Nat; ::_thesis: ( ( for i being Nat st i < j holds [(f . j),(f . i)] in the InternalRel of R ) implies for i being Nat st i < j + 1 holds [(f . (j + 1)),(f . i)] in the InternalRel of R ) assume A4: for i being Nat st i < j holds [(f . j),(f . i)] in the InternalRel of R ; ::_thesis: for i being Nat st i < j + 1 holds [(f . (j + 1)),(f . i)] in the InternalRel of R let i be Nat; ::_thesis: ( i < j + 1 implies [(f . (j + 1)),(f . i)] in the InternalRel of R ) assume A5: i < j + 1 ; ::_thesis: [(f . (j + 1)),(f . i)] in the InternalRel of R now__::_thesis:_[(f_._(j_+_1)),(f_._i)]_in_the_InternalRel_of_R percases ( i > j or i = j or i < j ) by XXREAL_0:1; suppose i > j ; ::_thesis: [(f . (j + 1)),(f . i)] in the InternalRel of R hence [(f . (j + 1)),(f . i)] in the InternalRel of R by A5, NAT_1:13; ::_thesis: verum end; suppose i = j ; ::_thesis: [(f . (j + 1)),(f . i)] in the InternalRel of R hence [(f . (j + 1)),(f . i)] in the InternalRel of R by A1, Def3; ::_thesis: verum end; suppose i < j ; ::_thesis: [(f . (j + 1)),(f . i)] in the InternalRel of R then A6: [(f . j),(f . i)] in the InternalRel of R by A4; [(f . (j + 1)),(f . j)] in the InternalRel of R by A1, Def3; hence [(f . (j + 1)),(f . i)] in the InternalRel of R by A2, A6, RELAT_2:def_8; ::_thesis: verum end; end; end; hence [(f . (j + 1)),(f . i)] in the InternalRel of R ; ::_thesis: verum end; then A7: for j being Nat st S1[j] holds S1[j + 1] ; thus for j being Nat holds S1[j] from NAT_1:sch_2(A3, A7); ::_thesis: verum end; theorem Th10: :: BAGORDER:10 for R being non empty transitive RelStr for s being sequence of R st R is well_founded & s is non-increasing holds ex p being Nat st for r being Nat st p <= r holds s . p = s . r proof let R be non empty transitive RelStr ; ::_thesis: for s being sequence of R st R is well_founded & s is non-increasing holds ex p being Nat st for r being Nat st p <= r holds s . p = s . r let s be sequence of R; ::_thesis: ( R is well_founded & s is non-increasing implies ex p being Nat st for r being Nat st p <= r holds s . p = s . r ) assume that A1: R is well_founded and A2: s is non-increasing ; ::_thesis: ex p being Nat st for r being Nat st p <= r holds s . p = s . r set cr = the carrier of R; set ir = the InternalRel of R; A3: the InternalRel of R is_well_founded_in the carrier of R by A1, WELLFND1:def_2; A4: dom s = NAT by FUNCT_2:def_1; rng s c= the carrier of R by RELAT_1:def_19; then consider a being set such that A5: a in rng s and A6: the InternalRel of R -Seg a misses rng s by A3, WELLORD1:def_3; A7: ( the InternalRel of R -Seg a) /\ (rng s) = {} by A6, XBOOLE_0:def_7; consider i being set such that A8: i in dom s and A9: s . i = a by A5, FUNCT_1:def_3; reconsider i = i as Nat by A8; assume for p being Nat ex r being Nat st ( p <= r & not s . p = s . r ) ; ::_thesis: contradiction then consider r being Nat such that A10: i <= r and A11: s . i <> s . r ; i < r by A10, A11, XXREAL_0:1; then [(s . r),(s . i)] in the InternalRel of R by A2, Th9; then A12: s . r in the InternalRel of R -Seg a by A9, A11, WELLORD1:1; reconsider r = r as Element of NAT by ORDINAL1:def_12; s . r in rng s by A4, FUNCT_1:3; hence contradiction by A7, A12, XBOOLE_0:def_4; ::_thesis: verum end; theorem Th11: :: BAGORDER:11 for X being set for a being Element of X for A being finite Subset of X for R being Order of X st A = {a} & R linearly_orders A holds SgmX (R,A) = <*a*> proof let X be set ; ::_thesis: for a being Element of X for A being finite Subset of X for R being Order of X st A = {a} & R linearly_orders A holds SgmX (R,A) = <*a*> let a be Element of X; ::_thesis: for A being finite Subset of X for R being Order of X st A = {a} & R linearly_orders A holds SgmX (R,A) = <*a*> let A be finite Subset of X; ::_thesis: for R being Order of X st A = {a} & R linearly_orders A holds SgmX (R,A) = <*a*> let R be Order of X; ::_thesis: ( A = {a} & R linearly_orders A implies SgmX (R,A) = <*a*> ) assume that A1: A = {a} and A2: R linearly_orders A ; ::_thesis: SgmX (R,A) = <*a*> A3: len (SgmX (R,A)) = card A by A2, PRE_POLY:11 .= 1 by A1, CARD_1:30 ; rng (SgmX (R,A)) = A by A2, PRE_POLY:def_2; hence SgmX (R,A) = <*a*> by A1, A3, FINSEQ_1:39; ::_thesis: verum end; begin definition let n be Ordinal; let b be bag of n; func TotDegree b -> Nat means :Def4: :: BAGORDER:def 4 ex f being FinSequence of NAT st ( it = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ); existence ex b1 being Nat ex f being FinSequence of NAT st ( b1 = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) proof set f = b * (SgmX ((RelIncl n),(support b))); A1: dom b = n by PARTFUN1:def_2; rng b c= NAT by VALUED_0:def_6; then reconsider bb = b as Function of n,NAT by A1, FUNCT_2:2; bb = b ; then reconsider f = b * (SgmX ((RelIncl n),(support b))) as FinSequence of NAT by FINSEQ_2:32; reconsider x = Sum f as Nat ; take x ; ::_thesis: ex f being FinSequence of NAT st ( x = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) thus ex f being FinSequence of NAT st ( x = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) ; ::_thesis: verum end; uniqueness for b1, b2 being Nat st ex f being FinSequence of NAT st ( b1 = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) & ex f being FinSequence of NAT st ( b2 = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) holds b1 = b2 ; end; :: deftheorem Def4 defines TotDegree BAGORDER:def_4_:_ for n being Ordinal for b being bag of n for b3 being Nat holds ( b3 = TotDegree b iff ex f being FinSequence of NAT st ( b3 = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) ); theorem Th12: :: BAGORDER:12 for n being Ordinal for b being bag of n for s being finite Subset of n for f, g being FinSequence of NAT st f = b * (SgmX ((RelIncl n),(support b))) & g = b * (SgmX ((RelIncl n),((support b) \/ s))) holds Sum f = Sum g proof let n be Ordinal; ::_thesis: for b being bag of n for s being finite Subset of n for f, g being FinSequence of NAT st f = b * (SgmX ((RelIncl n),(support b))) & g = b * (SgmX ((RelIncl n),((support b) \/ s))) holds Sum f = Sum g let b be bag of n; ::_thesis: for s being finite Subset of n for f, g being FinSequence of NAT st f = b * (SgmX ((RelIncl n),(support b))) & g = b * (SgmX ((RelIncl n),((support b) \/ s))) holds Sum f = Sum g let s be finite Subset of n; ::_thesis: for f, g being FinSequence of NAT st f = b * (SgmX ((RelIncl n),(support b))) & g = b * (SgmX ((RelIncl n),((support b) \/ s))) holds Sum f = Sum g let f, g be FinSequence of NAT ; ::_thesis: ( f = b * (SgmX ((RelIncl n),(support b))) & g = b * (SgmX ((RelIncl n),((support b) \/ s))) implies Sum f = Sum g ) assume that A1: f = b * (SgmX ((RelIncl n),(support b))) and A2: g = b * (SgmX ((RelIncl n),((support b) \/ s))) ; ::_thesis: Sum f = Sum g set sb = support b; set sbs = (support b) \/ s; set sbs9b = ((support b) \/ s) \ (support b); set xsb = SgmX ((RelIncl n),(support b)); set xsbs = SgmX ((RelIncl n),((support b) \/ s)); set xsbs9b = SgmX ((RelIncl n),(((support b) \/ s) \ (support b))); set xs = (SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b)))); set h = b * ((SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b))))); A3: dom b = n by PARTFUN1:def_2; A4: field (RelIncl n) = n by WELLORD2:def_1; A5: RelIncl n is being_linear-order by ORDERS_1:19; A6: RelIncl n linearly_orders n by A4, ORDERS_1:19, ORDERS_1:37; A7: RelIncl n linearly_orders (support b) \/ s by A4, A5, ORDERS_1:37, ORDERS_1:38; A8: RelIncl n linearly_orders support b by A4, A5, ORDERS_1:37, ORDERS_1:38; A9: RelIncl n linearly_orders ((support b) \/ s) \ (support b) by A4, A5, ORDERS_1:37, ORDERS_1:38; A10: rng (SgmX ((RelIncl n),((support b) \/ s))) = (support b) \/ s by A7, PRE_POLY:def_2; A11: rng (SgmX ((RelIncl n),(support b))) = support b by A8, PRE_POLY:def_2; A12: rng (SgmX ((RelIncl n),(((support b) \/ s) \ (support b)))) = ((support b) \/ s) \ (support b) by A9, PRE_POLY:def_2; then A13: rng ((SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b))))) = (support b) \/ (((support b) \/ s) \ (support b)) by A11, FINSEQ_1:31; then reconsider h = b * ((SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b))))) as FinSequence by A3, FINSEQ_1:16; percases ( n = {} or n <> {} ) ; suppose n = {} ; ::_thesis: Sum f = Sum g hence Sum f = Sum g by A1, A2; ::_thesis: verum end; suppose n <> {} ; ::_thesis: Sum f = Sum g then reconsider n = n as non empty Ordinal ; reconsider xsb = SgmX ((RelIncl n),(support b)), xsbs9b = SgmX ((RelIncl n),(((support b) \/ s) \ (support b))) as FinSequence of n ; rng b c= REAL ; then reconsider b = b as Function of n,REAL by A3, FUNCT_2:2; rng h c= rng b by RELAT_1:26; then rng h c= REAL by XBOOLE_1:1; then reconsider h = h as FinSequence of REAL by FINSEQ_1:def_4; reconsider gr = g as FinSequence of REAL by FINSEQ_2:24; A14: support b misses ((support b) \/ s) \ (support b) by XBOOLE_1:79; A15: (support b) \/ s = ((support b) \/ (support b)) \/ s .= (support b) \/ ((support b) \/ s) by XBOOLE_1:4 .= (support b) \/ (((support b) \/ s) \ (support b)) by XBOOLE_1:39 ; len ((SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b))))) = (len xsb) + (len xsbs9b) by FINSEQ_1:22 .= (card (support b)) + (len xsbs9b) by A6, ORDERS_1:38, PRE_POLY:11 .= (card (support b)) + (card (((support b) \/ s) \ (support b))) by A6, ORDERS_1:38, PRE_POLY:11 .= card ((support b) \/ s) by A15, CARD_2:40, XBOOLE_1:79 .= len (SgmX ((RelIncl n),((support b) \/ s))) by A6, ORDERS_1:38, PRE_POLY:11 ; then A16: dom (SgmX ((RelIncl n),((support b) \/ s))) = dom ((SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b))))) by FINSEQ_3:29; A17: SgmX ((RelIncl n),((support b) \/ s)) is one-to-one by A6, ORDERS_1:38, PRE_POLY:10; A18: rng xsb = support b by A8, PRE_POLY:def_2; A19: rng xsbs9b = ((support b) \/ s) \ (support b) by A9, PRE_POLY:def_2; A20: xsb is one-to-one by A6, ORDERS_1:38, PRE_POLY:10; xsbs9b is one-to-one by A6, ORDERS_1:38, PRE_POLY:10; then (SgmX ((RelIncl n),(support b))) ^ (SgmX ((RelIncl n),(((support b) \/ s) \ (support b)))) is one-to-one by A14, A18, A19, A20, FINSEQ_3:91; then A21: gr,h are_fiberwise_equipotent by A2, A3, A10, A13, A15, A16, A17, CLASSES1:83, RFINSEQ:26; now__::_thesis:_(_dom_xsbs9b_=_dom_(b_*_xsbs9b)_&_dom_xsbs9b_=_dom_((len_xsbs9b)_|->_0)_&_(_for_x_being_set_st_x_in_dom_xsbs9b_holds_ (b_*_xsbs9b)_._x_=_((len_xsbs9b)_|->_0)_._x_)_) thus dom xsbs9b = dom (b * xsbs9b) by A3, A12, RELAT_1:27; ::_thesis: ( dom xsbs9b = dom ((len xsbs9b) |-> 0) & ( for x being set st x in dom xsbs9b holds (b * xsbs9b) . x = ((len xsbs9b) |-> 0) . x ) ) A22: dom xsbs9b = Seg (len xsbs9b) by FINSEQ_1:def_3; hence dom xsbs9b = dom ((len xsbs9b) |-> 0) by FUNCOP_1:13; ::_thesis: for x being set st x in dom xsbs9b holds (b * xsbs9b) . x = ((len xsbs9b) |-> 0) . x let x be set ; ::_thesis: ( x in dom xsbs9b implies (b * xsbs9b) . x = ((len xsbs9b) |-> 0) . x ) assume A23: x in dom xsbs9b ; ::_thesis: (b * xsbs9b) . x = ((len xsbs9b) |-> 0) . x then xsbs9b . x in rng xsbs9b by FUNCT_1:3; then not xsbs9b . x in support b by A12, XBOOLE_0:def_5; then b . (xsbs9b . x) = 0 by PRE_POLY:def_7; hence (b * xsbs9b) . x = 0 by A23, FUNCT_1:13 .= ((len xsbs9b) |-> 0) . x by A22, A23, FUNCOP_1:7 ; ::_thesis: verum end; then A24: b * xsbs9b = (len xsbs9b) |-> 0 by FUNCT_1:2; h = (b * xsb) ^ (b * xsbs9b) by FINSEQOP:9; then Sum h = (Sum (b * xsb)) + (Sum (b * xsbs9b)) by RVSUM_1:75 .= (Sum f) + 0 by A1, A24, RVSUM_1:81 ; hence Sum f = Sum g by A21, RFINSEQ:9; ::_thesis: verum end; end; end; theorem Th13: :: BAGORDER:13 for n being Ordinal for a, b being bag of n holds TotDegree (a + b) = (TotDegree a) + (TotDegree b) proof let n be Ordinal; ::_thesis: for a, b being bag of n holds TotDegree (a + b) = (TotDegree a) + (TotDegree b) let a, b be bag of n; ::_thesis: TotDegree (a + b) = (TotDegree a) + (TotDegree b) A1: field (RelIncl n) = n by WELLORD2:def_1; A2: RelIncl n is being_linear-order by ORDERS_1:19; consider fab being FinSequence of NAT such that A3: TotDegree (a + b) = Sum fab and A4: fab = (a + b) * (SgmX ((RelIncl n),(support (a + b)))) by Def4; consider fa being FinSequence of NAT such that A5: TotDegree a = Sum fa and A6: fa = a * (SgmX ((RelIncl n),(support a))) by Def4; consider fb being FinSequence of NAT such that A7: TotDegree b = Sum fb and A8: fb = b * (SgmX ((RelIncl n),(support b))) by Def4; reconsider fab9 = fab as FinSequence of REAL by FINSEQ_2:24; set sasb = (support a) \/ (support b); reconsider sasb = (support a) \/ (support b) as finite Subset of n ; set s = SgmX ((RelIncl n),sasb); set fa9b = a * (SgmX ((RelIncl n),sasb)); set fb9a = b * (SgmX ((RelIncl n),sasb)); RelIncl n linearly_orders sasb by A1, A2, ORDERS_1:37, ORDERS_1:38; then A9: rng (SgmX ((RelIncl n),sasb)) = sasb by PRE_POLY:def_2; A10: support (a + b) = sasb by PRE_POLY:38; A11: dom a = n by PARTFUN1:def_2; A12: dom b = n by PARTFUN1:def_2; then reconsider fa9b = a * (SgmX ((RelIncl n),sasb)), fb9a = b * (SgmX ((RelIncl n),sasb)) as FinSequence by A9, A11, FINSEQ_1:16; A13: rng fa9b c= rng a by RELAT_1:26; A14: rng fb9a c= rng b by RELAT_1:26; A15: rng fa9b c= NAT by VALUED_0:def_6; A16: rng fb9a c= NAT by VALUED_0:def_6; A17: rng fa9b c= REAL by A13, XBOOLE_1:1; rng fb9a c= REAL by A14, XBOOLE_1:1; then reconsider fa9b = fa9b, fb9a = fb9a as FinSequence of REAL by A17, FINSEQ_1:def_4; reconsider fa9bn = fa9b, fb9an = fb9a as FinSequence of NAT by A15, A16, FINSEQ_1:def_4; set ln = len fab; A18: dom (a + b) = n by PARTFUN1:def_2; A19: Seg (len fab) = dom fab by FINSEQ_1:def_3 .= dom (SgmX ((RelIncl n),sasb)) by A4, A9, A10, A18, RELAT_1:27 ; then A20: Seg (len fab) = dom fa9b by A9, A11, RELAT_1:27; A21: Seg (len fab) = dom fb9a by A9, A12, A19, RELAT_1:27; A22: Sum fa = Sum fa9bn by A6, Th12; A23: Sum fb = Sum fb9an by A8, Th12; A24: len fa9b = len fb9a by A20, A21, FINSEQ_3:29; then A25: len (fa9b + fb9a) = len fa9b by INTEGRA5:2; then A26: Seg (len fab) = dom (fa9b + fb9a) by A20, FINSEQ_3:29; reconsider fa9b9 = fa9b as natural-valued ManySortedSet of Seg (len fab) by A20, PARTFUN1:def_2, RELAT_1:def_18; now__::_thesis:_(_Seg_(len_fab)_=_dom_fab9_&_Seg_(len_fab)_=_dom_(fa9b_+_fb9a)_&_(_for_k_being_Nat_st_k_in_Seg_(len_fab)_holds_ fab9_._k_=_(fa9b_+_fb9a)_._k_)_) thus Seg (len fab) = dom fab9 by FINSEQ_1:def_3; ::_thesis: ( Seg (len fab) = dom (fa9b + fb9a) & ( for k being Nat st k in Seg (len fab) holds fab9 . k = (fa9b + fb9a) . k ) ) thus Seg (len fab) = dom (fa9b + fb9a) by A20, A25, FINSEQ_3:29; ::_thesis: for k being Nat st k in Seg (len fab) holds fab9 . k = (fa9b + fb9a) . k let k be Nat; ::_thesis: ( k in Seg (len fab) implies fab9 . k = (fa9b + fb9a) . k ) assume A27: k in Seg (len fab) ; ::_thesis: fab9 . k = (fa9b + fb9a) . k reconsider k1 = k as Nat ; reconsider fa9bk = fa9b . k1, fb9ak = fb9a . k1 as Real ; thus fab9 . k = (a + b) . ((SgmX ((RelIncl n),(support (a + b)))) . k) by A4, A10, A19, A27, FUNCT_1:13 .= (a . ((SgmX ((RelIncl n),(support (a + b)))) . k)) + (b . ((SgmX ((RelIncl n),(support (a + b)))) . k)) by PRE_POLY:def_5 .= (fa9b9 . k) + (b . ((SgmX ((RelIncl n),(support (a + b)))) . k)) by A10, A19, A27, FUNCT_1:13 .= fa9bk + fb9ak by A10, A19, A27, FUNCT_1:13 .= (fa9b + fb9a) . k by A26, A27, VALUED_1:def_1 ; ::_thesis: verum end; then fab9 = fa9b + fb9a by FINSEQ_1:13; hence TotDegree (a + b) = (TotDegree a) + (TotDegree b) by A3, A5, A7, A22, A23, A24, INTEGRA5:2; ::_thesis: verum end; theorem :: BAGORDER:14 for n being Ordinal for a, b being bag of n st b divides a holds TotDegree (a -' b) = (TotDegree a) - (TotDegree b) proof let n be Ordinal; ::_thesis: for a, b being bag of n st b divides a holds TotDegree (a -' b) = (TotDegree a) - (TotDegree b) let a, b be bag of n; ::_thesis: ( b divides a implies TotDegree (a -' b) = (TotDegree a) - (TotDegree b) ) assume b divides a ; ::_thesis: TotDegree (a -' b) = (TotDegree a) - (TotDegree b) then A1: (a -' b) + b = a by PRE_POLY:47; TotDegree ((a -' b) + b) = (TotDegree (a -' b)) + (TotDegree b) by Th13; hence TotDegree (a -' b) = (TotDegree a) - (TotDegree b) by A1; ::_thesis: verum end; theorem Th15: :: BAGORDER:15 for n being Ordinal for b being bag of n holds ( TotDegree b = 0 iff b = EmptyBag n ) proof let n be Ordinal; ::_thesis: for b being bag of n holds ( TotDegree b = 0 iff b = EmptyBag n ) let b be bag of n; ::_thesis: ( TotDegree b = 0 iff b = EmptyBag n ) A1: field (RelIncl n) = n by WELLORD2:def_1; RelIncl n is being_linear-order by ORDERS_1:19; then A2: RelIncl n linearly_orders support b by A1, ORDERS_1:37, ORDERS_1:38; A3: dom b = n by PARTFUN1:def_2; hereby ::_thesis: ( b = EmptyBag n implies TotDegree b = 0 ) assume A4: TotDegree b = 0 ; ::_thesis: b = EmptyBag n consider f being FinSequence of NAT such that A5: TotDegree b = Sum f and A6: f = b * (SgmX ((RelIncl n),(support b))) by Def4; A7: f = (len f) |-> 0 by A4, A5, Th4; now__::_thesis:_for_z_being_set_st_z_in_dom_b_holds_ not_b_._z_<>_0 let z be set ; ::_thesis: ( z in dom b implies not b . z <> 0 ) assume that z in dom b and A8: b . z <> 0 ; ::_thesis: contradiction A9: rng (SgmX ((RelIncl n),(support b))) = support b by A2, PRE_POLY:def_2; z in support b by A8, PRE_POLY:def_7; then consider x being set such that A10: x in dom (SgmX ((RelIncl n),(support b))) and A11: (SgmX ((RelIncl n),(support b))) . x = z by A9, FUNCT_1:def_3; x in dom f by A3, A6, A9, A10, RELAT_1:27; then x in Seg (len f) by A7, FUNCOP_1:13; then f . x = 0 by A7, FUNCOP_1:7; hence contradiction by A6, A8, A10, A11, FUNCT_1:13; ::_thesis: verum end; then b = n --> 0 by A3, FUNCOP_1:11; hence b = EmptyBag n by PRE_POLY:def_13; ::_thesis: verum end; assume b = EmptyBag n ; ::_thesis: TotDegree b = 0 then A12: b = n --> 0 by PRE_POLY:def_13; A13: ex f being FinSequence of NAT st ( TotDegree b = Sum f & f = b * (SgmX ((RelIncl n),(support b))) ) by Def4; now__::_thesis:_not_support_b_<>_{} assume support b <> {} ; ::_thesis: contradiction then consider x being set such that A14: x in support b by XBOOLE_0:def_1; b . x = 0 by A12, A14, FUNCOP_1:7; hence contradiction by A14, PRE_POLY:def_7; ::_thesis: verum end; then rng (SgmX ((RelIncl n),(support b))) = {} by A2, PRE_POLY:def_2; then dom (SgmX ((RelIncl n),(support b))) = {} by RELAT_1:42; then dom (b * (SgmX ((RelIncl n),(support b)))) = {} by RELAT_1:25, XBOOLE_1:3; hence TotDegree b = 0 by A13, RELAT_1:41, RVSUM_1:72; ::_thesis: verum end; theorem Th16: :: BAGORDER:16 for i, j, n being Nat holds (i,j) -cut (EmptyBag n) = EmptyBag (j -' i) proof let i, j, n be Nat; ::_thesis: (i,j) -cut (EmptyBag n) = EmptyBag (j -' i) set CUT1 = (i,j) -cut (EmptyBag n); A1: dom ((i,j) -cut (EmptyBag n)) = j -' i by PARTFUN1:def_2; now__::_thesis:_for_k_being_set_holds_((i,j)_-cut_(EmptyBag_n))_._k_<=_(EmptyBag_(j_-'_i))_._k let k be set ; ::_thesis: ((i,j) -cut (EmptyBag n)) . b1 <= (EmptyBag (j -' i)) . b1 percases ( k in dom ((i,j) -cut (EmptyBag n)) or not k in dom ((i,j) -cut (EmptyBag n)) ) ; supposeA2: k in dom ((i,j) -cut (EmptyBag n)) ; ::_thesis: ((i,j) -cut (EmptyBag n)) . b1 <= (EmptyBag (j -' i)) . b1 j -' i = { x where x is Element of NAT : x < j -' i } by AXIOMS:4; then ex x being Element of NAT st ( k = x & x < j -' i ) by A1, A2; then reconsider k9 = k as Element of NAT ; ((i,j) -cut (EmptyBag n)) . k = (EmptyBag n) . (i + k9) by A2, Def1 .= 0 by PRE_POLY:52 ; hence ((i,j) -cut (EmptyBag n)) . k <= (EmptyBag (j -' i)) . k ; ::_thesis: verum end; suppose not k in dom ((i,j) -cut (EmptyBag n)) ; ::_thesis: ((i,j) -cut (EmptyBag n)) . b1 <= (EmptyBag (j -' i)) . b1 hence ((i,j) -cut (EmptyBag n)) . k <= (EmptyBag (j -' i)) . k by FUNCT_1:def_2; ::_thesis: verum end; end; end; then (i,j) -cut (EmptyBag n) divides EmptyBag (j -' i) by PRE_POLY:def_11; hence (i,j) -cut (EmptyBag n) = EmptyBag (j -' i) by PRE_POLY:58; ::_thesis: verum end; theorem Th17: :: BAGORDER:17 for i, j, n being Nat for a, b being bag of n holds (i,j) -cut (a + b) = ((i,j) -cut a) + ((i,j) -cut b) proof let i, j, n be Nat; ::_thesis: for a, b being bag of n holds (i,j) -cut (a + b) = ((i,j) -cut a) + ((i,j) -cut b) let a, b be bag of n; ::_thesis: (i,j) -cut (a + b) = ((i,j) -cut a) + ((i,j) -cut b) set CUTAB = (i,j) -cut (a + b); set CUTA = (i,j) -cut a; set CUTB = (i,j) -cut b; now__::_thesis:_for_x_being_set_st_x_in_j_-'_i_holds_ ((i,j)_-cut_(a_+_b))_._x_=_(((i,j)_-cut_a)_+_((i,j)_-cut_b))_._x let x be set ; ::_thesis: ( x in j -' i implies ((i,j) -cut (a + b)) . x = (((i,j) -cut a) + ((i,j) -cut b)) . x ) assume A1: x in j -' i ; ::_thesis: ((i,j) -cut (a + b)) . x = (((i,j) -cut a) + ((i,j) -cut b)) . x j -' i = { k where k is Element of NAT : k < j -' i } by AXIOMS:4; then ex k being Element of NAT st ( k = x & k < j -' i ) by A1; then reconsider x9 = x as Element of NAT ; ((i,j) -cut (a + b)) . x = (a + b) . (i + x9) by A1, Def1; then A2: ((i,j) -cut (a + b)) . x = (a . (i + x9)) + (b . (i + x9)) by PRE_POLY:def_5; A3: ((i,j) -cut a) . x = a . (i + x9) by A1, Def1; ((i,j) -cut b) . x = b . (i + x9) by A1, Def1; hence ((i,j) -cut (a + b)) . x = (((i,j) -cut a) + ((i,j) -cut b)) . x by A2, A3, PRE_POLY:def_5; ::_thesis: verum end; hence (i,j) -cut (a + b) = ((i,j) -cut a) + ((i,j) -cut b) by PBOOLE:3; ::_thesis: verum end; theorem :: BAGORDER:18 for X being set holds support (EmptyBag X) = {} proof let n be set ; ::_thesis: support (EmptyBag n) = {} assume not support (EmptyBag n) = {} ; ::_thesis: contradiction then consider x being set such that A1: x in support (EmptyBag n) by XBOOLE_0:def_1; (EmptyBag n) . x <> 0 by A1, PRE_POLY:def_7; hence contradiction by PRE_POLY:52; ::_thesis: verum end; theorem Th19: :: BAGORDER:19 for X being set for b being bag of X st support b = {} holds b = EmptyBag X proof let n be set ; ::_thesis: for b being bag of n st support b = {} holds b = EmptyBag n let b be bag of n; ::_thesis: ( support b = {} implies b = EmptyBag n ) assume that A1: support b = {} and A2: b <> EmptyBag n ; ::_thesis: contradiction A3: dom b = n by PARTFUN1:def_2; dom (EmptyBag n) = n by PARTFUN1:def_2; then consider x being set such that x in n and A4: b . x <> (EmptyBag n) . x by A2, A3, FUNCT_1:2; b . x <> 0 by A4, PRE_POLY:52; hence contradiction by A1, PRE_POLY:def_7; ::_thesis: verum end; theorem Th20: :: BAGORDER:20 for n, m being Ordinal for b being bag of n st m in n holds b | m is bag of m proof let n, m be Ordinal; ::_thesis: for b being bag of n st m in n holds b | m is bag of m let b be bag of n; ::_thesis: ( m in n implies b | m is bag of m ) assume m in n ; ::_thesis: b | m is bag of m then A1: m c= n by ORDINAL1:def_2; dom b = n by PARTFUN1:def_2; then dom (b | m) = m by A1, RELAT_1:62; hence b | m is bag of m by PARTFUN1:def_2; ::_thesis: verum end; theorem :: BAGORDER:21 for n being Ordinal for a, b being bag of n st b divides a holds support b c= support a proof let n be Ordinal; ::_thesis: for a, b being bag of n st b divides a holds support b c= support a let a, b be bag of n; ::_thesis: ( b divides a implies support b c= support a ) assume A1: b divides a ; ::_thesis: support b c= support a let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in support b or x in support a ) assume x in support b ; ::_thesis: x in support a then b . x <> 0 by PRE_POLY:def_7; then a . x <> 0 by A1, PRE_POLY:def_11; hence x in support a by PRE_POLY:def_7; ::_thesis: verum end; begin definition let n be set ; mode TermOrder of n is Order of (Bags n); end; notation let n be Ordinal; synonym LexOrder n for BagOrder n; end; definition let n be Ordinal; let T be TermOrder of n; attrT is admissible means :Def5: :: BAGORDER:def 5 ( T is_strongly_connected_in Bags n & ( for a being bag of n holds [(EmptyBag n),a] in T ) & ( for a, b, c being bag of n st [a,b] in T holds [(a + c),(b + c)] in T ) ); end; :: deftheorem Def5 defines admissible BAGORDER:def_5_:_ for n being Ordinal for T being TermOrder of n holds ( T is admissible iff ( T is_strongly_connected_in Bags n & ( for a being bag of n holds [(EmptyBag n),a] in T ) & ( for a, b, c being bag of n st [a,b] in T holds [(a + c),(b + c)] in T ) ) ); theorem Th22: :: BAGORDER:22 for n being Ordinal holds LexOrder n is admissible proof let n be Ordinal; ::_thesis: LexOrder n is admissible now__::_thesis:_for_a,_b_being_set_st_a_in_Bags_n_&_b_in_Bags_n_&_not_[a,b]_in_BagOrder_n_holds_ [b,a]_in_BagOrder_n let a, b be set ; ::_thesis: ( a in Bags n & b in Bags n & not [a,b] in BagOrder n implies [b,a] in BagOrder n ) assume that A1: a in Bags n and A2: b in Bags n ; ::_thesis: ( [a,b] in BagOrder n or [b,a] in BagOrder n ) reconsider a9 = a, b9 = b as bag of n by A1, A2; ( a9 <=' b9 or b9 <=' a9 ) by PRE_POLY:45; hence ( [a,b] in BagOrder n or [b,a] in BagOrder n ) by PRE_POLY:def_14; ::_thesis: verum end; hence LexOrder n is_strongly_connected_in Bags n by RELAT_2:def_7; :: according to BAGORDER:def_5 ::_thesis: ( ( for a being bag of n holds [(EmptyBag n),a] in LexOrder n ) & ( for a, b, c being bag of n st [a,b] in LexOrder n holds [(a + c),(b + c)] in LexOrder n ) ) now__::_thesis:_for_a_being_bag_of_n_holds_[(EmptyBag_n),a]_in_BagOrder_n let a be bag of n; ::_thesis: [(EmptyBag n),a] in BagOrder n EmptyBag n <=' a by PRE_POLY:60; hence [(EmptyBag n),a] in BagOrder n by PRE_POLY:def_14; ::_thesis: verum end; hence for a being bag of n holds [(EmptyBag n),a] in LexOrder n ; ::_thesis: for a, b, c being bag of n st [a,b] in LexOrder n holds [(a + c),(b + c)] in LexOrder n now__::_thesis:_for_a,_b,_c_being_bag_of_n_st_[a,b]_in_BagOrder_n_holds_ [(a_+_c),(b_+_c)]_in_BagOrder_n let a, b, c be bag of n; ::_thesis: ( [a,b] in BagOrder n implies [(a + c),(b + c)] in BagOrder n ) assume [a,b] in BagOrder n ; ::_thesis: [(a + c),(b + c)] in BagOrder n then A3: a <=' b by PRE_POLY:def_14; now__::_thesis:_a_+_c_<='_b_+_c percases ( a < b or a = b ) by A3, PRE_POLY:def_10; suppose a < b ; ::_thesis: a + c <=' b + c then consider k being Ordinal such that A4: a . k < b . k and A5: for l being Ordinal st l in k holds a . l = b . l by PRE_POLY:def_9; now__::_thesis:_ex_k_being_Ordinal_st_ (_(a_+_c)_._k_<_(b_+_c)_._k_&_(_for_l_being_Ordinal_st_l_in_k_holds_ (a_+_c)_._l_=_(b_+_c)_._l_)_) take k = k; ::_thesis: ( (a + c) . k < (b + c) . k & ( for l being Ordinal st l in k holds (a + c) . l = (b + c) . l ) ) A6: (a + c) . k = (a . k) + (c . k) by PRE_POLY:def_5; (b + c) . k = (b . k) + (c . k) by PRE_POLY:def_5; hence (a + c) . k < (b + c) . k by A4, A6, XREAL_1:6; ::_thesis: for l being Ordinal st l in k holds (a + c) . l = (b + c) . l let l be Ordinal; ::_thesis: ( l in k implies (a + c) . l = (b + c) . l ) assume A7: l in k ; ::_thesis: (a + c) . l = (b + c) . l A8: (a + c) . l = (a . l) + (c . l) by PRE_POLY:def_5; (b + c) . l = (b . l) + (c . l) by PRE_POLY:def_5; hence (a + c) . l = (b + c) . l by A5, A7, A8; ::_thesis: verum end; then a + c < b + c by PRE_POLY:def_9; hence a + c <=' b + c by PRE_POLY:def_10; ::_thesis: verum end; suppose a = b ; ::_thesis: a + c <=' b + c hence a + c <=' b + c ; ::_thesis: verum end; end; end; hence [(a + c),(b + c)] in BagOrder n by PRE_POLY:def_14; ::_thesis: verum end; hence for a, b, c being bag of n st [a,b] in LexOrder n holds [(a + c),(b + c)] in LexOrder n ; ::_thesis: verum end; registration let n be Ordinal; cluster Relation-like Bags n -defined Bags n -valued total V18( Bags n, Bags n) reflexive antisymmetric transitive admissible for Element of bool [:(Bags n),(Bags n):]; existence ex b1 being TermOrder of n st b1 is admissible proof LexOrder n is admissible by Th22; hence ex b1 being TermOrder of n st b1 is admissible ; ::_thesis: verum end; end; registration let n be Ordinal; cluster LexOrder n -> admissible ; coherence LexOrder n is admissible by Th22; end; theorem :: BAGORDER:23 for o being infinite Ordinal holds not LexOrder o is well-ordering proof let o be infinite Ordinal; ::_thesis: not LexOrder o is well-ordering set R = LexOrder o; set r = RelStr(# (Bags o),(LexOrder o) #); set ir = the InternalRel of RelStr(# (Bags o),(LexOrder o) #); set cr = the carrier of RelStr(# (Bags o),(LexOrder o) #); assume LexOrder o is well-ordering ; ::_thesis: contradiction then A1: LexOrder o is well_founded ; the carrier of RelStr(# (Bags o),(LexOrder o) #) = field the InternalRel of RelStr(# (Bags o),(LexOrder o) #) by ORDERS_1:15; then the InternalRel of RelStr(# (Bags o),(LexOrder o) #) is_well_founded_in the carrier of RelStr(# (Bags o),(LexOrder o) #) by A1, WELLORD1:3; then A2: RelStr(# (Bags o),(LexOrder o) #) is well_founded by WELLFND1:def_2; defpred S1[ set , set ] means $2 = (o --> 0) +* ($1,1); A3: now__::_thesis:_for_n_being_Element_of_NAT_ex_y_being_Element_of_the_carrier_of_RelStr(#_(Bags_o),(LexOrder_o)_#)_st_S1[n,y] let n be Element of NAT ; ::_thesis: ex y being Element of the carrier of RelStr(# (Bags o),(LexOrder o) #) st S1[n,y] set y = (o --> 0) +* (n,1); A4: dom (o --> 0) = o by FUNCOP_1:13; reconsider y = (o --> 0) +* (n,1) as ManySortedSet of o ; A5: n in omega ; A6: omega c= o by CARD_3:85; now__::_thesis:_for_x_being_set_holds_ (_(_x_in_{n}_implies_y_._x_<>_0_)_&_(_y_._x_<>_0_implies_x_in_{n}_)_) let x be set ; ::_thesis: ( ( x in {n} implies y . x <> 0 ) & ( y . x <> 0 implies x in {n} ) ) hereby ::_thesis: ( y . x <> 0 implies x in {n} ) assume x in {n} ; ::_thesis: y . x <> 0 then x = n by TARSKI:def_1; hence y . x <> 0 by A4, A5, A6, FUNCT_7:31; ::_thesis: verum end; assume that A7: y . x <> 0 and A8: not x in {n} ; ::_thesis: contradiction x <> n by A8, TARSKI:def_1; then A9: y . x = (o --> 0) . x by FUNCT_7:32; percases ( x in dom (o --> 0) or not x in dom (o --> 0) ) ; suppose x in dom (o --> 0) ; ::_thesis: contradiction hence contradiction by A7, A9, FUNCOP_1:7; ::_thesis: verum end; suppose not x in dom (o --> 0) ; ::_thesis: contradiction hence contradiction by A7, A9, FUNCT_1:def_2; ::_thesis: verum end; end; end; then support y = {n} by PRE_POLY:def_7; then y is finite-support by PRE_POLY:def_8; then reconsider y = y as Element of the carrier of RelStr(# (Bags o),(LexOrder o) #) by PRE_POLY:def_12; take y = y; ::_thesis: S1[n,y] thus S1[n,y] ; ::_thesis: verum end; consider f being Function of NAT, the carrier of RelStr(# (Bags o),(LexOrder o) #) such that A10: for n being Element of NAT holds S1[n,f . n] from FUNCT_2:sch_3(A3); reconsider f = f as sequence of RelStr(# (Bags o),(LexOrder o) #) ; f is descending proof let n be Nat; :: according to WELLFND1:def_6 ::_thesis: ( not f . (n + 1) = f . n & [(f . (n + 1)),(f . n)] in the InternalRel of RelStr(# (Bags o),(LexOrder o) #) ) reconsider n0 = n as Element of NAT by ORDINAL1:def_12; set fn1 = f . (n0 + 1); set fn = f . n0; A11: f . (n0 + 1) = (o --> 0) +* ((n + 1),1) by A10; A12: f . n0 = (o --> 0) +* (n,1) by A10; reconsider fn1 = f . (n0 + 1) as bag of o ; reconsider fn = f . n0 as bag of o ; A13: n0 in omega ; A14: omega c= o by CARD_3:85; n <> n + 1 ; then A15: fn1 . n = (o --> 0) . n by A11, FUNCT_7:32 .= 0 by A13, A14, FUNCOP_1:7 ; A16: dom (o --> 0) = o by FUNCOP_1:13; then A17: fn . n = 1 by A12, A13, A14, FUNCT_7:31; now__::_thesis:_for_l_being_Ordinal_st_l_in_n_holds_ fn1_._l_=_fn_._l let l be Ordinal; ::_thesis: ( l in n implies fn1 . l = fn . l ) assume A18: l in n ; ::_thesis: fn1 . l = fn . l then A19: l <> n ; n < n + 1 by NAT_1:13; then n in { i where i is Element of NAT : i < n0 + 1 } ; then n in n + 1 by AXIOMS:4; then n c= n + 1 by ORDINAL1:def_2; then l in n + 1 by A18; then l <> n + 1 ; hence fn1 . l = (o --> 0) . l by A11, FUNCT_7:32 .= fn . l by A12, A19, FUNCT_7:32 ; ::_thesis: verum end; then A20: fn1 < fn by A15, A17, PRE_POLY:def_9; thus f . (n + 1) <> f . n by A12, A13, A14, A15, A16, FUNCT_7:31; ::_thesis: [(f . (n + 1)),(f . n)] in the InternalRel of RelStr(# (Bags o),(LexOrder o) #) fn1 <=' fn by A20, PRE_POLY:def_10; hence [(f . (n + 1)),(f . n)] in the InternalRel of RelStr(# (Bags o),(LexOrder o) #) by PRE_POLY:def_14; ::_thesis: verum end; hence contradiction by A2, WELLFND1:14; ::_thesis: verum end; definition let n be Ordinal; func InvLexOrder n -> TermOrder of n means :Def6: :: BAGORDER:def 6 for p, q being bag of n holds ( [p,q] in it iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ); existence ex b1 being TermOrder of n st for p, q being bag of n holds ( [p,q] in b1 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) proof defpred S1[ set , set ] means ( $1 = $2 or ex p, q being Element of Bags n st ( $1 = p & $2 = q & ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ); consider ILO being Relation of (Bags n),(Bags n) such that A1: for x, y being set holds ( [x,y] in ILO iff ( x in Bags n & y in Bags n & S1[x,y] ) ) from RELSET_1:sch_1(); A2: ILO is_reflexive_in Bags n proof let x be set ; :: according to RELAT_2:def_1 ::_thesis: ( not x in Bags n or [x,x] in ILO ) assume x in Bags n ; ::_thesis: [x,x] in ILO hence [x,x] in ILO by A1; ::_thesis: verum end; A3: ILO is_antisymmetric_in Bags n proof let x, y be set ; :: according to RELAT_2:def_4 ::_thesis: ( not x in Bags n or not y in Bags n or not [x,y] in ILO or not [y,x] in ILO or x = y ) assume that x in Bags n and y in Bags n and A4: [x,y] in ILO and A5: [y,x] in ILO ; ::_thesis: x = y percases ( x = y or not x = y ) ; suppose x = y ; ::_thesis: x = y hence x = y ; ::_thesis: verum end; supposeA6: not x = y ; ::_thesis: x = y then consider p, q being Element of Bags n such that A7: x = p and A8: y = q and A9: ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) by A1, A4; ex q9, p9 being Element of Bags n st ( y = q9 & x = p9 & ex i being Ordinal st ( i in n & q9 . i < p9 . i & ( for k being Ordinal st i in k & k in n holds q9 . k = p9 . k ) ) ) by A1, A5, A6; then consider i being Ordinal such that A10: i in n and A11: q . i < p . i and A12: for k being Ordinal st i in k & k in n holds q . k = p . k by A7, A8; consider j being Ordinal such that A13: j in n and A14: p . j < q . j and A15: for k being Ordinal st j in k & k in n holds p . k = q . k by A9; now__::_thesis:_i_=_j percases ( i in j or i = j or j in i ) by ORDINAL1:14; suppose i in j ; ::_thesis: i = j hence i = j by A12, A13, A14; ::_thesis: verum end; suppose i = j ; ::_thesis: i = j hence i = j ; ::_thesis: verum end; suppose j in i ; ::_thesis: i = j hence i = j by A10, A11, A15; ::_thesis: verum end; end; end; hence x = y by A11, A14; ::_thesis: verum end; end; end; A16: ILO is_transitive_in Bags n proof let x, y, z be set ; :: according to RELAT_2:def_8 ::_thesis: ( not x in Bags n or not y in Bags n or not z in Bags n or not [x,y] in ILO or not [y,z] in ILO or [x,z] in ILO ) assume that x in Bags n and y in Bags n and z in Bags n and A17: [x,y] in ILO and A18: [y,z] in ILO ; ::_thesis: [x,z] in ILO percases ( x = y or x <> y ) ; suppose x = y ; ::_thesis: [x,z] in ILO hence [x,z] in ILO by A18; ::_thesis: verum end; suppose x <> y ; ::_thesis: [x,z] in ILO then consider p, q being Element of Bags n such that A19: x = p and A20: y = q and A21: ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) by A1, A17; consider i being Ordinal such that A22: i in n and A23: p . i < q . i and A24: for k being Ordinal st i in k & k in n holds p . k = q . k by A21; now__::_thesis:_[x,z]_in_ILO percases ( y = z or y <> z ) ; suppose y = z ; ::_thesis: [x,z] in ILO hence [x,z] in ILO by A17; ::_thesis: verum end; suppose y <> z ; ::_thesis: [x,z] in ILO then consider q9, r being Element of Bags n such that A25: y = q9 and A26: z = r and A27: ex i being Ordinal st ( i in n & q9 . i < r . i & ( for k being Ordinal st i in k & k in n holds q9 . k = r . k ) ) by A1, A18; consider j being Ordinal such that A28: j in n and A29: q9 . j < r . j and A30: for k being Ordinal st j in k & k in n holds q9 . k = r . k by A27; now__::_thesis:_[x,z]_in_ILO percases ( i in j or i = j or j in i ) by ORDINAL1:14; supposeA31: i in j ; ::_thesis: [x,z] in ILO then A32: p . j < r . j by A20, A24, A25, A28, A29; now__::_thesis:_for_k_being_Ordinal_st_j_in_k_&_k_in_n_holds_ p_._k_=_r_._k let k be Ordinal; ::_thesis: ( j in k & k in n implies p . k = r . k ) assume that A33: j in k and A34: k in n ; ::_thesis: p . k = r . k A35: q . k = r . k by A20, A25, A30, A33, A34; i in k by A31, A33, ORDINAL1:10; hence p . k = r . k by A24, A34, A35; ::_thesis: verum end; hence [x,z] in ILO by A1, A19, A26, A28, A32; ::_thesis: verum end; supposeA36: i = j ; ::_thesis: [x,z] in ILO now__::_thesis:_ex_p,_r_being_Element_of_Bags_n_st_ (_x_=_p_&_z_=_r_&_ex_j_being_Ordinal_st_ (_j_in_n_&_p_._j_<_r_._j_&_(_for_k_being_Ordinal_st_j_in_k_&_k_in_n_holds_ p_._k_=_r_._k_)_)_) take p = p; ::_thesis: ex r being Element of Bags n st ( x = p & z = r & ex j being Ordinal st ( j in n & p . j < r . j & ( for k being Ordinal st j in k & k in n holds p . k = r . k ) ) ) take r = r; ::_thesis: ( x = p & z = r & ex j being Ordinal st ( j in n & p . j < r . j & ( for k being Ordinal st j in k & k in n holds p . k = r . k ) ) ) thus ( x = p & z = r ) by A19, A26; ::_thesis: ex j being Ordinal st ( j in n & p . j < r . j & ( for k being Ordinal st j in k & k in n holds p . k = r . k ) ) take j = j; ::_thesis: ( j in n & p . j < r . j & ( for k being Ordinal st j in k & k in n holds p . k = r . k ) ) thus j in n by A28; ::_thesis: ( p . j < r . j & ( for k being Ordinal st j in k & k in n holds p . k = r . k ) ) thus p . j < r . j by A20, A23, A25, A29, A36, XXREAL_0:2; ::_thesis: for k being Ordinal st j in k & k in n holds p . k = r . k now__::_thesis:_for_k_being_Ordinal_st_j_in_k_&_k_in_n_holds_ p_._k_=_r_._k let k be Ordinal; ::_thesis: ( j in k & k in n implies p . k = r . k ) assume that A37: j in k and A38: k in n ; ::_thesis: p . k = r . k p . k = q . k by A24, A36, A37, A38; hence p . k = r . k by A20, A25, A30, A37, A38; ::_thesis: verum end; hence for k being Ordinal st j in k & k in n holds p . k = r . k ; ::_thesis: verum end; hence [x,z] in ILO by A1; ::_thesis: verum end; supposeA39: j in i ; ::_thesis: [x,z] in ILO then A40: p . i < r . i by A20, A22, A23, A25, A30; now__::_thesis:_for_k_being_Ordinal_st_i_in_k_&_k_in_n_holds_ p_._k_=_r_._k let k be Ordinal; ::_thesis: ( i in k & k in n implies p . k = r . k ) assume that A41: i in k and A42: k in n ; ::_thesis: p . k = r . k A43: p . k = q . k by A24, A41, A42; j in k by A39, A41, ORDINAL1:10; hence p . k = r . k by A20, A25, A30, A42, A43; ::_thesis: verum end; hence [x,z] in ILO by A1, A19, A22, A26, A40; ::_thesis: verum end; end; end; hence [x,z] in ILO ; ::_thesis: verum end; end; end; hence [x,z] in ILO ; ::_thesis: verum end; end; end; A44: dom ILO = Bags n by A2, ORDERS_1:13; field ILO = Bags n by A2, ORDERS_1:13; then reconsider ILO = ILO as TermOrder of n by A2, A3, A16, A44, PARTFUN1:def_2, RELAT_2:def_9, RELAT_2:def_12, RELAT_2:def_16; take ILO ; ::_thesis: for p, q being bag of n holds ( [p,q] in ILO iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) let x, y be bag of n; ::_thesis: ( [x,y] in ILO iff ( x = y or ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) ) ) hereby ::_thesis: ( ( x = y or ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) ) implies [x,y] in ILO ) assume A45: [x,y] in ILO ; ::_thesis: ( x = y or ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) ) now__::_thesis:_(_x_<>_y_implies_ex_i_being_Ordinal_st_ (_i_in_n_&_x_._i_<_y_._i_&_(_for_k_being_Ordinal_st_i_in_k_&_k_in_n_holds_ x_._k_=_y_._k_)_)_) assume x <> y ; ::_thesis: ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) then ex p, q being Element of Bags n st ( x = p & y = q & ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) by A1, A45; hence ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) ; ::_thesis: verum end; hence ( x = y or ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) ) ; ::_thesis: verum end; assume A46: ( x = y or ex i being Ordinal st ( i in n & x . i < y . i & ( for k being Ordinal st i in k & k in n holds x . k = y . k ) ) ) ; ::_thesis: [x,y] in ILO A47: now__::_thesis:_(_x_<>_y_implies_ex_p,_q_being_Element_of_Bags_n_st_ (_x_=_p_&_y_=_q_&_ex_i_being_Ordinal_st_ (_i_in_n_&_p_._i_<_q_._i_&_(_for_k_being_Ordinal_st_i_in_k_&_k_in_n_holds_ p_._k_=_q_._k_)_)_)_) assume A48: x <> y ; ::_thesis: ex p, q being Element of Bags n st ( x = p & y = q & ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) thus ex p, q being Element of Bags n st ( x = p & y = q & ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ::_thesis: verum proof reconsider x9 = x, y9 = y as Element of Bags n by PRE_POLY:def_12; take x9 ; ::_thesis: ex q being Element of Bags n st ( x = x9 & y = q & ex i being Ordinal st ( i in n & x9 . i < q . i & ( for k being Ordinal st i in k & k in n holds x9 . k = q . k ) ) ) take y9 ; ::_thesis: ( x = x9 & y = y9 & ex i being Ordinal st ( i in n & x9 . i < y9 . i & ( for k being Ordinal st i in k & k in n holds x9 . k = y9 . k ) ) ) thus ( x = x9 & y = y9 ) ; ::_thesis: ex i being Ordinal st ( i in n & x9 . i < y9 . i & ( for k being Ordinal st i in k & k in n holds x9 . k = y9 . k ) ) thus ex i being Ordinal st ( i in n & x9 . i < y9 . i & ( for k being Ordinal st i in k & k in n holds x9 . k = y9 . k ) ) by A46, A48; ::_thesis: verum end; end; x in Bags n by PRE_POLY:def_12; hence [x,y] in ILO by A1, A47; ::_thesis: verum end; uniqueness for b1, b2 being TermOrder of n st ( for p, q being bag of n holds ( [p,q] in b1 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) ) & ( for p, q being bag of n holds ( [p,q] in b2 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) ) holds b1 = b2 proof let IT1, IT2 be TermOrder of n; ::_thesis: ( ( for p, q being bag of n holds ( [p,q] in IT1 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) ) & ( for p, q being bag of n holds ( [p,q] in IT2 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) ) implies IT1 = IT2 ) assume that A49: for p, q being bag of n holds ( [p,q] in IT1 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) and A50: for p, q being bag of n holds ( [p,q] in IT2 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_IT1_implies_[a,b]_in_IT2_)_&_(_[a,b]_in_IT2_implies_[a,b]_in_IT1_)_) let a, b be set ; ::_thesis: ( ( [a,b] in IT1 implies [a,b] in IT2 ) & ( [a,b] in IT2 implies [b1,b2] in IT1 ) ) hereby ::_thesis: ( [a,b] in IT2 implies [b1,b2] in IT1 ) assume A51: [a,b] in IT1 ; ::_thesis: [a,b] in IT2 then consider p, q being set such that A52: [a,b] = [p,q] and A53: p in Bags n and A54: q in Bags n by RELSET_1:2; reconsider p = p, q = q as bag of n by A53, A54; percases ( p = q or p <> q ) ; suppose p = q ; ::_thesis: [a,b] in IT2 hence [a,b] in IT2 by A50, A52; ::_thesis: verum end; suppose p <> q ; ::_thesis: [a,b] in IT2 then ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) by A49, A51, A52; hence [a,b] in IT2 by A50, A52; ::_thesis: verum end; end; end; assume A55: [a,b] in IT2 ; ::_thesis: [b1,b2] in IT1 then consider p, q being set such that A56: [a,b] = [p,q] and A57: p in Bags n and A58: q in Bags n by RELSET_1:2; reconsider p = p, q = q as bag of n by A57, A58; percases ( p = q or p <> q ) ; suppose p = q ; ::_thesis: [b1,b2] in IT1 hence [a,b] in IT1 by A49, A56; ::_thesis: verum end; suppose p <> q ; ::_thesis: [b1,b2] in IT1 then ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) by A50, A55, A56; hence [a,b] in IT1 by A49, A56; ::_thesis: verum end; end; end; hence IT1 = IT2 by RELAT_1:def_2; ::_thesis: verum end; end; :: deftheorem Def6 defines InvLexOrder BAGORDER:def_6_:_ for n being Ordinal for b2 being TermOrder of n holds ( b2 = InvLexOrder n iff for p, q being bag of n holds ( [p,q] in b2 iff ( p = q or ex i being Ordinal st ( i in n & p . i < q . i & ( for k being Ordinal st i in k & k in n holds p . k = q . k ) ) ) ) ); theorem Th24: :: BAGORDER:24 for n being Ordinal holds InvLexOrder n is admissible proof let n be Ordinal; ::_thesis: InvLexOrder n is admissible set ILO = InvLexOrder n; now__::_thesis:_for_x,_y_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_not_[x,y]_in_InvLexOrder_n_holds_ [y,x]_in_InvLexOrder_n let x, y be set ; ::_thesis: ( x in Bags n & y in Bags n & not [x,y] in InvLexOrder n implies [y,x] in InvLexOrder n ) assume that A1: x in Bags n and A2: y in Bags n ; ::_thesis: ( [x,y] in InvLexOrder n or [y,x] in InvLexOrder n ) reconsider p = x, q = y as bag of n by A1, A2; now__::_thesis:_(_not_[p,q]_in_InvLexOrder_n_implies_[q,p]_in_InvLexOrder_n_) assume A3: not [p,q] in InvLexOrder n ; ::_thesis: [q,p] in InvLexOrder n then A4: p <> q by Def6; set s = SgmX ((RelIncl n),((support p) \/ (support q))); A5: dom p = n by PARTFUN1:def_2; A6: dom q = n by PARTFUN1:def_2; A7: field (RelIncl n) = n by WELLORD2:def_1; RelIncl n is being_linear-order by ORDERS_1:19; then A8: RelIncl n linearly_orders (support p) \/ (support q) by A7, ORDERS_1:37, ORDERS_1:38; then A9: rng (SgmX ((RelIncl n),((support p) \/ (support q)))) = (support p) \/ (support q) by PRE_POLY:def_2; defpred S1[ Nat] means ( $1 in dom (SgmX ((RelIncl n),((support p) \/ (support q)))) & q . ((SgmX ((RelIncl n),((support p) \/ (support q)))) . $1) <> p . ((SgmX ((RelIncl n),((support p) \/ (support q)))) . $1) ); A10: for k being Nat st S1[k] holds k <= len (SgmX ((RelIncl n),((support p) \/ (support q)))) by FINSEQ_3:25; A11: ex k being Nat st S1[k] proof assume A12: for k being Nat holds not S1[k] ; ::_thesis: contradiction now__::_thesis:_for_x_being_set_st_x_in_n_holds_ p_._x_=_q_._x let x be set ; ::_thesis: ( x in n implies p . b1 = q . b1 ) assume x in n ; ::_thesis: p . b1 = q . b1 percases ( p . x <> 0 or q . x <> 0 or ( p . x = 0 & q . x = 0 ) ) ; suppose p . x <> 0 ; ::_thesis: p . b1 = q . b1 then x in support p by PRE_POLY:def_7; then x in (support p) \/ (support q) by XBOOLE_0:def_3; then ex k being Nat st ( k in dom (SgmX ((RelIncl n),((support p) \/ (support q)))) & (SgmX ((RelIncl n),((support p) \/ (support q)))) . k = x ) by A9, FINSEQ_2:10; hence p . x = q . x by A12; ::_thesis: verum end; suppose q . x <> 0 ; ::_thesis: p . b1 = q . b1 then x in support q by PRE_POLY:def_7; then x in (support p) \/ (support q) by XBOOLE_0:def_3; then ex k being Nat st ( k in dom (SgmX ((RelIncl n),((support p) \/ (support q)))) & (SgmX ((RelIncl n),((support p) \/ (support q)))) . k = x ) by A9, FINSEQ_2:10; hence p . x = q . x by A12; ::_thesis: verum end; suppose ( p . x = 0 & q . x = 0 ) ; ::_thesis: p . b1 = q . b1 hence p . x = q . x ; ::_thesis: verum end; end; end; hence contradiction by A4, A5, A6, FUNCT_1:2; ::_thesis: verum end; consider j being Nat such that A13: S1[j] and A14: for k being Nat st S1[k] holds k <= j from NAT_1:sch_6(A10, A11); A15: (SgmX ((RelIncl n),((support p) \/ (support q)))) . j in rng (SgmX ((RelIncl n),((support p) \/ (support q)))) by A13, FUNCT_1:3; then reconsider J = (SgmX ((RelIncl n),((support p) \/ (support q)))) . j as Ordinal by A9; now__::_thesis:_ex_J_being_Ordinal_st_ (_J_in_n_&_q_._J_<_p_._J_&_(_for_k_being_Ordinal_st_J_in_k_&_k_in_n_holds_ q_._k_=_p_._k_)_) take J = J; ::_thesis: ( J in n & q . J < p . J & ( for k being Ordinal st J in k & k in n holds q . k = p . k ) ) thus J in n by A9, A15; ::_thesis: ( q . J < p . J & ( for k being Ordinal st J in k & k in n holds q . k = p . k ) ) A16: now__::_thesis:_for_k_being_Ordinal_st_J_in_k_&_k_in_n_holds_ not_q_._k_<>_p_._k let k be Ordinal; ::_thesis: ( J in k & k in n implies not q . k <> p . k ) assume that A17: J in k and A18: k in n and A19: q . k <> p . k ; ::_thesis: contradiction now__::_thesis:_(_not_k_in_support_p_implies_k_in_support_q_) assume that A20: not k in support p and A21: not k in support q ; ::_thesis: contradiction p . k = 0 by A20, PRE_POLY:def_7; hence contradiction by A19, A21, PRE_POLY:def_7; ::_thesis: verum end; then k in (support p) \/ (support q) by XBOOLE_0:def_3; then consider m being Nat such that A22: m in dom (SgmX ((RelIncl n),((support p) \/ (support q)))) and A23: (SgmX ((RelIncl n),((support p) \/ (support q)))) . m = k by A9, FINSEQ_2:10; A24: m <= j by A14, A19, A22, A23; m <> j by A17, A23; then m < j by A24, XXREAL_0:1; then [((SgmX ((RelIncl n),((support p) \/ (support q)))) /. m),((SgmX ((RelIncl n),((support p) \/ (support q)))) /. j)] in RelIncl n by A8, A13, A22, PRE_POLY:def_2; then [((SgmX ((RelIncl n),((support p) \/ (support q)))) . m),((SgmX ((RelIncl n),((support p) \/ (support q)))) /. j)] in RelIncl n by A22, PARTFUN1:def_6; then [((SgmX ((RelIncl n),((support p) \/ (support q)))) . m),((SgmX ((RelIncl n),((support p) \/ (support q)))) . j)] in RelIncl n by A13, PARTFUN1:def_6; then (SgmX ((RelIncl n),((support p) \/ (support q)))) . m c= (SgmX ((RelIncl n),((support p) \/ (support q)))) . j by A9, A15, A18, A23, WELLORD2:def_1; hence contradiction by A17, A23, ORDINAL1:5; ::_thesis: verum end; then q . J <= p . J by A3, A9, A15, Def6; hence q . J < p . J by A13, XXREAL_0:1; ::_thesis: for k being Ordinal st J in k & k in n holds q . k = p . k thus for k being Ordinal st J in k & k in n holds q . k = p . k by A16; ::_thesis: verum end; hence [q,p] in InvLexOrder n by Def6; ::_thesis: verum end; hence ( [x,y] in InvLexOrder n or [y,x] in InvLexOrder n ) ; ::_thesis: verum end; hence InvLexOrder n is_strongly_connected_in Bags n by RELAT_2:def_7; :: according to BAGORDER:def_5 ::_thesis: ( ( for a being bag of n holds [(EmptyBag n),a] in InvLexOrder n ) & ( for a, b, c being bag of n st [a,b] in InvLexOrder n holds [(a + c),(b + c)] in InvLexOrder n ) ) now__::_thesis:_for_a_being_bag_of_n_holds_[(EmptyBag_n),a]_in_InvLexOrder_n let a be bag of n; ::_thesis: [(EmptyBag n),b1] in InvLexOrder n percases ( EmptyBag n = a or EmptyBag n <> a ) ; suppose EmptyBag n = a ; ::_thesis: [(EmptyBag n),b1] in InvLexOrder n hence [(EmptyBag n),a] in InvLexOrder n by Def6; ::_thesis: verum end; supposeA25: EmptyBag n <> a ; ::_thesis: [(EmptyBag n),b1] in InvLexOrder n set s = SgmX ((RelIncl n),(support a)); A26: field (RelIncl n) = n by WELLORD2:def_1; RelIncl n is being_linear-order by ORDERS_1:19; then A27: RelIncl n linearly_orders support a by A26, ORDERS_1:37, ORDERS_1:38; then A28: rng (SgmX ((RelIncl n),(support a))) = support a by PRE_POLY:def_2; then rng (SgmX ((RelIncl n),(support a))) <> {} by A25, Th19; then A29: len (SgmX ((RelIncl n),(support a))) in dom (SgmX ((RelIncl n),(support a))) by FINSEQ_5:6, RELAT_1:38; then A30: (SgmX ((RelIncl n),(support a))) . (len (SgmX ((RelIncl n),(support a)))) in rng (SgmX ((RelIncl n),(support a))) by FUNCT_1:3; then reconsider j = (SgmX ((RelIncl n),(support a))) . (len (SgmX ((RelIncl n),(support a)))) as Ordinal by A28; now__::_thesis:_ex_j_being_Ordinal_st_ (_j_in_n_&_(EmptyBag_n)_._j_<_a_._j_&_(_for_k_being_Ordinal_st_j_in_k_&_k_in_n_holds_ (EmptyBag_n)_._k_=_a_._k_)_) take j = j; ::_thesis: ( j in n & (EmptyBag n) . j < a . j & ( for k being Ordinal st j in k & k in n holds (EmptyBag n) . k = a . k ) ) thus j in n by A28, A30; ::_thesis: ( (EmptyBag n) . j < a . j & ( for k being Ordinal st j in k & k in n holds (EmptyBag n) . k = a . k ) ) A31: a . j <> 0 by A28, A30, PRE_POLY:def_7; (EmptyBag n) . j = 0 by PRE_POLY:52; hence (EmptyBag n) . j < a . j by A31; ::_thesis: for k being Ordinal st j in k & k in n holds (EmptyBag n) . k = a . k let k be Ordinal; ::_thesis: ( j in k & k in n implies (EmptyBag n) . k = a . k ) assume that A32: j in k and A33: k in n ; ::_thesis: (EmptyBag n) . k = a . k A34: j c= k by A32, ORDINAL1:def_2; now__::_thesis:_not_(EmptyBag_n)_._k_<>_a_._k assume (EmptyBag n) . k <> a . k ; ::_thesis: contradiction then a . k <> 0 by PRE_POLY:52; then k in support a by PRE_POLY:def_7; then consider i being Nat such that A35: i in dom (SgmX ((RelIncl n),(support a))) and A36: (SgmX ((RelIncl n),(support a))) . i = k by A28, FINSEQ_2:10; A37: i <= len (SgmX ((RelIncl n),(support a))) by A35, FINSEQ_3:25; percases ( i = len (SgmX ((RelIncl n),(support a))) or i < len (SgmX ((RelIncl n),(support a))) ) by A37, XXREAL_0:1; suppose i = len (SgmX ((RelIncl n),(support a))) ; ::_thesis: contradiction hence contradiction by A32, A36; ::_thesis: verum end; suppose i < len (SgmX ((RelIncl n),(support a))) ; ::_thesis: contradiction then [((SgmX ((RelIncl n),(support a))) /. i),((SgmX ((RelIncl n),(support a))) /. (len (SgmX ((RelIncl n),(support a)))))] in RelIncl n by A27, A29, A35, PRE_POLY:def_2; then [((SgmX ((RelIncl n),(support a))) . i),((SgmX ((RelIncl n),(support a))) /. (len (SgmX ((RelIncl n),(support a)))))] in RelIncl n by A35, PARTFUN1:def_6; then [((SgmX ((RelIncl n),(support a))) . i),((SgmX ((RelIncl n),(support a))) . (len (SgmX ((RelIncl n),(support a)))))] in RelIncl n by A29, PARTFUN1:def_6; then k c= j by A28, A30, A33, A36, WELLORD2:def_1; then j = k by A34, XBOOLE_0:def_10; hence contradiction by A32; ::_thesis: verum end; end; end; hence (EmptyBag n) . k = a . k ; ::_thesis: verum end; hence [(EmptyBag n),a] in InvLexOrder n by Def6; ::_thesis: verum end; end; end; hence for a being bag of n holds [(EmptyBag n),a] in InvLexOrder n ; ::_thesis: for a, b, c being bag of n st [a,b] in InvLexOrder n holds [(a + c),(b + c)] in InvLexOrder n now__::_thesis:_for_a,_b,_c_being_bag_of_n_st_[a,b]_in_InvLexOrder_n_holds_ [(a_+_c),(b_+_c)]_in_InvLexOrder_n let a, b, c be bag of n; ::_thesis: ( [a,b] in InvLexOrder n implies [(b1 + b3),(b2 + b3)] in InvLexOrder n ) assume A38: [a,b] in InvLexOrder n ; ::_thesis: [(b1 + b3),(b2 + b3)] in InvLexOrder n percases ( a = b or a <> b ) ; supposeA39: a = b ; ::_thesis: [(b1 + b3),(b2 + b3)] in InvLexOrder n a + c in Bags n by PRE_POLY:def_12; hence [(a + c),(b + c)] in InvLexOrder n by A39, ORDERS_1:3; ::_thesis: verum end; suppose a <> b ; ::_thesis: [(b1 + b3),(b2 + b3)] in InvLexOrder n then consider i being Ordinal such that A40: i in n and A41: a . i < b . i and A42: for k being Ordinal st i in k & k in n holds a . k = b . k by A38, Def6; now__::_thesis:_ex_i_being_Ordinal_st_ (_i_in_n_&_(a_+_c)_._i_<_(b_+_c)_._i_&_(_for_k_being_Ordinal_st_i_in_k_&_k_in_n_holds_ (a_+_c)_._k_=_(b_+_c)_._k_)_) take i = i; ::_thesis: ( i in n & (a + c) . i < (b + c) . i & ( for k being Ordinal st i in k & k in n holds (a + c) . k = (b + c) . k ) ) thus i in n by A40; ::_thesis: ( (a + c) . i < (b + c) . i & ( for k being Ordinal st i in k & k in n holds (a + c) . k = (b + c) . k ) ) A43: (a + c) . i = (a . i) + (c . i) by PRE_POLY:def_5; (b + c) . i = (b . i) + (c . i) by PRE_POLY:def_5; hence (a + c) . i < (b + c) . i by A41, A43, XREAL_1:6; ::_thesis: for k being Ordinal st i in k & k in n holds (a + c) . k = (b + c) . k let k be Ordinal; ::_thesis: ( i in k & k in n implies (a + c) . k = (b + c) . k ) assume that A44: i in k and A45: k in n ; ::_thesis: (a + c) . k = (b + c) . k A46: (a + c) . k = (a . k) + (c . k) by PRE_POLY:def_5; (b + c) . k = (b . k) + (c . k) by PRE_POLY:def_5; hence (a + c) . k = (b + c) . k by A42, A44, A45, A46; ::_thesis: verum end; hence [(a + c),(b + c)] in InvLexOrder n by Def6; ::_thesis: verum end; end; end; hence for a, b, c being bag of n st [a,b] in InvLexOrder n holds [(a + c),(b + c)] in InvLexOrder n ; ::_thesis: verum end; registration let n be Ordinal; cluster InvLexOrder n -> admissible ; coherence InvLexOrder n is admissible by Th24; end; theorem Th25: :: BAGORDER:25 for o being Ordinal holds InvLexOrder o is well-ordering proof defpred S1[ Ordinal] means InvLexOrder $1 is well-ordering ; A1: now__::_thesis:_for_o_being_Ordinal_st_(_for_n_being_Ordinal_st_n_in_o_holds_ S1[n]_)_holds_ S1[o] let o be Ordinal; ::_thesis: ( ( for n being Ordinal st n in o holds S1[n] ) implies S1[o] ) assume A2: for n being Ordinal st n in o holds S1[n] ; ::_thesis: S1[o] set ilo = InvLexOrder o; A3: InvLexOrder o is_strongly_connected_in Bags o by Def5; then InvLexOrder o is_reflexive_in Bags o by ORDERS_1:7; then A4: field (InvLexOrder o) = Bags o by PARTIT_2:9; A5: now__::_thesis:_InvLexOrder_o_is_well_founded assume not InvLexOrder o is well_founded ; ::_thesis: contradiction then A6: not InvLexOrder o is_well_founded_in Bags o by A4, WELLORD1:3; set R = RelStr(# (Bags o),(InvLexOrder o) #); set ir = the InternalRel of RelStr(# (Bags o),(InvLexOrder o) #); not RelStr(# (Bags o),(InvLexOrder o) #) is well_founded by A6, WELLFND1:def_2; then consider f being sequence of RelStr(# (Bags o),(InvLexOrder o) #) such that A7: f is descending by WELLFND1:14; reconsider a = f . 0 as bag of o ; set s = SgmX ((RelIncl o),(support a)); A8: field (RelIncl o) = o by WELLORD2:def_1; RelIncl o is being_linear-order by ORDERS_1:19; then A9: RelIncl o linearly_orders support a by A8, ORDERS_1:37, ORDERS_1:38; then A10: rng (SgmX ((RelIncl o),(support a))) = support a by PRE_POLY:def_2; now__::_thesis:_not_rng_(SgmX_((RelIncl_o),(support_a)))_=_{} assume rng (SgmX ((RelIncl o),(support a))) = {} ; ::_thesis: contradiction then A11: a = EmptyBag o by A10, Th19; reconsider b = f . (0 + 1) as bag of o ; A12: b <> a by A7, WELLFND1:def_6; [b,a] in the InternalRel of RelStr(# (Bags o),(InvLexOrder o) #) by A7, WELLFND1:def_6; then ex i being Ordinal st ( i in o & b . i < a . i & ( for k being Ordinal st i in k & k in o holds b . k = a . k ) ) by A12, Def6; hence contradiction by A11, PRE_POLY:52; ::_thesis: verum end; then A13: len (SgmX ((RelIncl o),(support a))) in dom (SgmX ((RelIncl o),(support a))) by FINSEQ_5:6, RELAT_1:38; then A14: (SgmX ((RelIncl o),(support a))) . (len (SgmX ((RelIncl o),(support a)))) in rng (SgmX ((RelIncl o),(support a))) by FUNCT_1:3; then reconsider j = (SgmX ((RelIncl o),(support a))) . (len (SgmX ((RelIncl o),(support a)))) as Ordinal by A10; defpred S2[ set , set ] means ex b being bag of o st ( f . $1 = b & $2 = b . j ); A15: now__::_thesis:_for_x_being_Element_of_NAT_ex_y_being_Element_of_NAT_st_S2[x,y] let x be Element of NAT ; ::_thesis: ex y being Element of NAT st S2[x,y] reconsider b = f . x as bag of o ; take y = b . j; ::_thesis: S2[x,y] thus S2[x,y] ; ::_thesis: verum end; consider t being Function of NAT,NAT such that A16: for i being Element of NAT holds S2[i,t . i] from FUNCT_2:sch_3(A15); defpred S3[ Nat] means for i being Ordinal for b being bag of o st j in i & i in o & f . $1 = b holds b . i = 0 ; A17: S3[ 0 ] proof let i be Ordinal; ::_thesis: for b being bag of o st j in i & i in o & f . 0 = b holds b . i = 0 let b be bag of o; ::_thesis: ( j in i & i in o & f . 0 = b implies b . i = 0 ) assume that A18: j in i and A19: i in o and A20: f . 0 = b ; ::_thesis: b . i = 0 assume b . i <> 0 ; ::_thesis: contradiction then i in support a by A20, PRE_POLY:def_7; then consider k being Nat such that A21: k in dom (SgmX ((RelIncl o),(support a))) and A22: (SgmX ((RelIncl o),(support a))) . k = i by A10, FINSEQ_2:10; A23: k <= len (SgmX ((RelIncl o),(support a))) by A21, FINSEQ_3:25; percases ( k = len (SgmX ((RelIncl o),(support a))) or k < len (SgmX ((RelIncl o),(support a))) ) by A23, XXREAL_0:1; suppose k = len (SgmX ((RelIncl o),(support a))) ; ::_thesis: contradiction hence contradiction by A18, A22; ::_thesis: verum end; suppose k < len (SgmX ((RelIncl o),(support a))) ; ::_thesis: contradiction then [((SgmX ((RelIncl o),(support a))) /. k),((SgmX ((RelIncl o),(support a))) /. (len (SgmX ((RelIncl o),(support a)))))] in RelIncl o by A9, A13, A21, PRE_POLY:def_2; then [((SgmX ((RelIncl o),(support a))) . k),((SgmX ((RelIncl o),(support a))) /. (len (SgmX ((RelIncl o),(support a)))))] in RelIncl o by A21, PARTFUN1:def_6; then [((SgmX ((RelIncl o),(support a))) . k),((SgmX ((RelIncl o),(support a))) . (len (SgmX ((RelIncl o),(support a)))))] in RelIncl o by A13, PARTFUN1:def_6; then (SgmX ((RelIncl o),(support a))) . k c= (SgmX ((RelIncl o),(support a))) . (len (SgmX ((RelIncl o),(support a)))) by A10, A14, A19, A22, WELLORD2:def_1; hence contradiction by A18, A22, ORDINAL1:5; ::_thesis: verum end; end; end; A24: for n being Nat st S3[n] holds S3[n + 1] proof let n be Nat; ::_thesis: ( S3[n] implies S3[n + 1] ) assume A25: for i being Ordinal for b being bag of o st j in i & i in o & f . n = b holds b . i = 0 ; ::_thesis: S3[n + 1] let i be Ordinal; ::_thesis: for b being bag of o st j in i & i in o & f . (n + 1) = b holds b . i = 0 let b1 be bag of o; ::_thesis: ( j in i & i in o & f . (n + 1) = b1 implies b1 . i = 0 ) assume that A26: j in i and A27: i in o and A28: f . (n + 1) = b1 ; ::_thesis: b1 . i = 0 reconsider b = f . n as bag of o ; A29: b . i = 0 by A25, A26, A27; A30: b1 <> b by A7, A28, WELLFND1:def_6; [b1,b] in InvLexOrder o by A7, A28, WELLFND1:def_6; then consider i1 being Ordinal such that A31: i1 in o and A32: b1 . i1 < b . i1 and A33: for k being Ordinal st i1 in k & k in o holds b1 . k = b . k by A30, Def6; percases ( i1 in i or i1 = i or i in i1 ) by ORDINAL1:14; suppose i1 in i ; ::_thesis: b1 . i = 0 hence b1 . i = 0 by A27, A29, A33; ::_thesis: verum end; suppose i1 = i ; ::_thesis: b1 . i = 0 hence b1 . i = 0 by A25, A26, A27, A32; ::_thesis: verum end; supposeA34: i in i1 ; ::_thesis: b1 . i = 0 assume b1 . i <> 0 ; ::_thesis: contradiction j in i1 by A26, A34, ORDINAL1:10; hence contradiction by A25, A31, A32; ::_thesis: verum end; end; end; A35: for n being Nat holds S3[n] from NAT_1:sch_2(A17, A24); reconsider t = t as sequence of OrderedNAT by DICKSON:def_15; A36: t is non-increasing proof let n be Nat; :: according to BAGORDER:def_3 ::_thesis: [(t . (n + 1)),(t . n)] in the InternalRel of OrderedNAT reconsider n0 = n as Element of NAT by ORDINAL1:def_12; reconsider tn = t . n0, tn1 = t . (n0 + 1) as Nat ; reconsider fn = f . n0, fn1 = f . (n0 + 1) as bag of o ; A37: fn1 <> fn by A7, WELLFND1:def_6; [fn1,fn] in InvLexOrder o by A7, WELLFND1:def_6; then consider i being Ordinal such that A38: i in o and A39: fn1 . i < fn . i and A40: for k being Ordinal st i in k & k in o holds fn1 . k = fn . k by A37, Def6; A41: ex bn being bag of o st ( fn = bn & tn = bn . j ) by A16; A42: ex bn1 being bag of o st ( fn1 = bn1 & tn1 = bn1 . j ) by A16; now__::_thesis:_tn1_<=_tn percases ( i = j or j in i or i in j ) by ORDINAL1:14; suppose i = j ; ::_thesis: tn1 <= tn hence tn1 <= tn by A39, A41, A42; ::_thesis: verum end; suppose j in i ; ::_thesis: tn1 <= tn hence tn1 <= tn by A35, A38, A39; ::_thesis: verum end; suppose i in j ; ::_thesis: tn1 <= tn hence tn1 <= tn by A10, A14, A40, A41, A42; ::_thesis: verum end; end; end; hence [(t . (n + 1)),(t . n)] in the InternalRel of OrderedNAT by DICKSON:def_14, DICKSON:def_15; ::_thesis: verum end; set n = j; set iln = InvLexOrder j; set S = RelStr(# (Bags j),(InvLexOrder j) #); InvLexOrder j is_strongly_connected_in Bags j by Def5; then InvLexOrder j is_reflexive_in Bags j by ORDERS_1:7; then A43: field (InvLexOrder j) = Bags j by PARTIT_2:9; consider p being Nat such that A44: for r being Nat st p <= r holds t . p = t . r by A36, Th10; defpred S4[ Nat, set ] means ex b being bag of o st ( b = f . (p + $1) & $2 = b | j ); A45: now__::_thesis:_for_x_being_Element_of_NAT_ex_y_being_Element_of_Bags_j_st_S4[x,y] let x be Element of NAT ; ::_thesis: ex y being Element of Bags j st S4[x,y] reconsider b = f . (p + x) as bag of o ; reconsider y = b | j as bag of j by A10, A14, Th20; reconsider y = y as Element of Bags j by PRE_POLY:def_12; take y = y; ::_thesis: S4[x,y] thus S4[x,y] ; ::_thesis: verum end; consider g being Function of NAT,(Bags j) such that A46: for x being Element of NAT holds S4[x,g . x] from FUNCT_2:sch_3(A45); reconsider g = g as sequence of RelStr(# (Bags j),(InvLexOrder j) #) ; now__::_thesis:_for_k_being_Nat_holds_ (_g_._(k_+_1)_<>_g_._k_&_[(g_._(k_+_1)),(g_._k)]_in_InvLexOrder_j_) let k be Nat; ::_thesis: ( g . (k + 1) <> g . k & [(g . (k + 1)),(g . k)] in InvLexOrder j ) reconsider k0 = k as Element of NAT by ORDINAL1:def_12; consider b being bag of o such that A47: b = f . (p + k) and A48: g . k0 = b | j by A46; consider b1 being bag of o such that A49: b1 = f . (p + (k + 1)) and A50: g . (k + 1) = b1 | j by A46; p + (k + 1) = (p + k) + 1 ; then A51: b <> b1 by A7, A47, A49, WELLFND1:def_6; A52: ex bb being bag of o st ( f . (p + k) = bb & t . (p + k0) = bb . j ) by A16; A53: ex bb1 being bag of o st ( f . (p + (k + 1)) = bb1 & t . (p + (k + 1)) = bb1 . j ) by A16; A54: t . (p + k) = t . p by A44, NAT_1:11; thus g . (k + 1) <> g . k ::_thesis: [(g . (k + 1)),(g . k)] in InvLexOrder j proof assume A55: not g . (k + 1) <> g . k ; ::_thesis: contradiction A56: o = dom b by PARTFUN1:def_2; A57: o = dom b1 by PARTFUN1:def_2; now__::_thesis:_for_m_being_set_st_m_in_o_holds_ b_._m_=_b1_._m let m be set ; ::_thesis: ( m in o implies b . b1 = b1 . b1 ) assume A58: m in o ; ::_thesis: b . b1 = b1 . b1 percases ( m in j or m = j or j in m ) by A58, ORDINAL1:14; supposeA59: m in j ; ::_thesis: b . b1 = b1 . b1 then (b | j) . m = b . m by FUNCT_1:49; hence b . m = b1 . m by A48, A50, A55, A59, FUNCT_1:49; ::_thesis: verum end; suppose m = j ; ::_thesis: b . b1 = b1 . b1 hence b . m = b1 . m by A44, A47, A49, A52, A53, A54, NAT_1:11; ::_thesis: verum end; supposeA60: j in m ; ::_thesis: b . b1 = b1 . b1 then b . m = 0 by A35, A47, A58; hence b . m = b1 . m by A35, A49, A58, A60; ::_thesis: verum end; end; end; hence contradiction by A51, A56, A57, FUNCT_1:2; ::_thesis: verum end; [(f . ((p + k) + 1)),(f . (p + k))] in InvLexOrder o by A7, WELLFND1:def_6; then consider i being Ordinal such that A61: i in o and A62: b1 . i < b . i and A63: for k being Ordinal st i in k & k in o holds b . k = b1 . k by A47, A49, A51, Def6; A64: now__::_thesis:_i_in_j assume A65: not i in j ; ::_thesis: contradiction percases ( i = j or j in i ) by A65, ORDINAL1:14; suppose i = j ; ::_thesis: contradiction hence contradiction by A44, A47, A49, A52, A53, A54, A62, NAT_1:11; ::_thesis: verum end; supposeA66: j in i ; ::_thesis: contradiction then b . i = 0 by A35, A47, A61 .= b1 . i by A35, A49, A61, A66 ; hence contradiction by A62; ::_thesis: verum end; end; end; reconsider bj = b | j, b1j = b1 | j as bag of j by A10, A14, Th20; A67: b1j . i = b1 . i by A64, FUNCT_1:49; A68: bj . i = b . i by A64, FUNCT_1:49; now__::_thesis:_for_k_being_Ordinal_st_i_in_k_&_k_in_j_holds_ bj_._k_=_b1j_._k let k be Ordinal; ::_thesis: ( i in k & k in j implies bj . k = b1j . k ) assume that A69: i in k and A70: k in j ; ::_thesis: bj . k = b1j . k k in o by A10, A14, A70, ORDINAL1:10; then A71: b . k = b1 . k by A63, A69; thus bj . k = b . k by A70, FUNCT_1:49 .= b1j . k by A70, A71, FUNCT_1:49 ; ::_thesis: verum end; hence [(g . (k + 1)),(g . k)] in InvLexOrder j by A48, A50, A62, A64, A67, A68, Def6; ::_thesis: verum end; then g is descending by WELLFND1:def_6; then not RelStr(# (Bags j),(InvLexOrder j) #) is well_founded by WELLFND1:14; then not InvLexOrder j is_well_founded_in the carrier of RelStr(# (Bags j),(InvLexOrder j) #) by WELLFND1:def_2; then not InvLexOrder j well_orders field (InvLexOrder j) by A43, WELLORD1:def_5; then not InvLexOrder j is well-ordering by WELLORD1:4; hence contradiction by A2, A10, A14; ::_thesis: verum end; A72: field (InvLexOrder o) = Bags o by ORDERS_1:12; then A73: InvLexOrder o is_reflexive_in Bags o by RELAT_2:def_9; A74: InvLexOrder o is_transitive_in Bags o by A72, RELAT_2:def_16; A75: InvLexOrder o is_antisymmetric_in Bags o by A72, RELAT_2:def_12; A76: InvLexOrder o is_connected_in field (InvLexOrder o) by A3, A4, ORDERS_1:7; InvLexOrder o is_well_founded_in field (InvLexOrder o) by A5, WELLORD1:3; then InvLexOrder o well_orders field (InvLexOrder o) by A4, A73, A74, A75, A76, WELLORD1:def_5; hence S1[o] by WELLORD1:4; ::_thesis: verum end; thus for o being Ordinal holds S1[o] from ORDINAL1:sch_2(A1); ::_thesis: verum end; definition let n be Ordinal; let o be TermOrder of n; assume B1: for a, b, c being bag of n st [a,b] in o holds [(a + c),(b + c)] in o ; func Graded o -> TermOrder of n means :Def7: :: BAGORDER:def 7 for a, b being bag of n holds ( [a,b] in it iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ); existence ex b1 being TermOrder of n st for a, b being bag of n holds ( [a,b] in b1 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) proof defpred S1[ set , set ] means ex x9, y9 being bag of n st ( x9 = $1 & y9 = $2 & ( TotDegree x9 < TotDegree y9 or ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ) ); consider R being Relation of (Bags n) such that A1: for x, y being set holds ( [x,y] in R iff ( x in Bags n & y in Bags n & S1[x,y] ) ) from RELSET_1:sch_1(); A2: now__::_thesis:_for_x_being_set_st_x_in_Bags_n_holds_ [x,x]_in_R let x be set ; ::_thesis: ( x in Bags n implies [x,x] in R ) assume A3: x in Bags n ; ::_thesis: [x,x] in R reconsider x9 = x as bag of n by A3; now__::_thesis:_ex_x9_being_bag_of_n_st_ (_x9_=_x_&_TotDegree_x9_=_TotDegree_x9_&_[x9,x9]_in_o_) take x9 = x9; ::_thesis: ( x9 = x & TotDegree x9 = TotDegree x9 & [x9,x9] in o ) thus x9 = x ; ::_thesis: ( TotDegree x9 = TotDegree x9 & [x9,x9] in o ) thus TotDegree x9 = TotDegree x9 ; ::_thesis: [x9,x9] in o [(EmptyBag n),(EmptyBag n)] in o by ORDERS_1:3; then [((EmptyBag n) + x9),((EmptyBag n) + x9)] in o by B1; then [x9,((EmptyBag n) + x9)] in o by PRE_POLY:53; hence [x9,x9] in o by PRE_POLY:53; ::_thesis: verum end; hence [x,x] in R by A1, A3; ::_thesis: verum end; A4: now__::_thesis:_for_x,_y_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_[x,y]_in_R_&_[y,x]_in_R_holds_ x_=_y let x, y be set ; ::_thesis: ( x in Bags n & y in Bags n & [x,y] in R & [y,x] in R implies x = y ) assume that A5: x in Bags n and A6: y in Bags n and A7: [x,y] in R and A8: [y,x] in R ; ::_thesis: x = y consider x9, y9 being bag of n such that A9: x9 = x and A10: y9 = y and A11: ( TotDegree x9 < TotDegree y9 or ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ) by A1, A7; A12: ex y99, x99 being bag of n st ( y99 = y & x99 = x & ( TotDegree y99 < TotDegree x99 or ( TotDegree y99 = TotDegree x99 & [y99,x99] in o ) ) ) by A1, A8; now__::_thesis:_(_TotDegree_x9_=_TotDegree_y9_&_[x9,y9]_in_o_&_TotDegree_y9_=_TotDegree_x9_&_[y9,x9]_in_o_) percases ( TotDegree x9 < TotDegree y9 or ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ) by A11; supposeA13: TotDegree x9 < TotDegree y9 ; ::_thesis: ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) now__::_thesis:_contradiction percases ( TotDegree y9 < TotDegree x9 or ( TotDegree y9 = TotDegree x9 & [y9,x9] in o ) ) by A9, A10, A12; suppose TotDegree y9 < TotDegree x9 ; ::_thesis: contradiction hence contradiction by A13; ::_thesis: verum end; suppose ( TotDegree y9 = TotDegree x9 & [y9,x9] in o ) ; ::_thesis: contradiction hence contradiction by A13; ::_thesis: verum end; end; end; hence ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) ; ::_thesis: verum end; supposeA14: ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ; ::_thesis: ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) now__::_thesis:_(_TotDegree_x9_=_TotDegree_y9_&_[x9,y9]_in_o_&_TotDegree_y9_=_TotDegree_x9_&_[y9,x9]_in_o_) percases ( TotDegree y9 < TotDegree x9 or ( TotDegree y9 = TotDegree x9 & [y9,x9] in o ) ) by A9, A10, A12; suppose TotDegree y9 < TotDegree x9 ; ::_thesis: ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) hence ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) by A14; ::_thesis: verum end; suppose ( TotDegree y9 = TotDegree x9 & [y9,x9] in o ) ; ::_thesis: ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) hence ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) by A14; ::_thesis: verum end; end; end; hence ( TotDegree x9 = TotDegree y9 & [x9,y9] in o & TotDegree y9 = TotDegree x9 & [y9,x9] in o ) ; ::_thesis: verum end; end; end; hence x = y by A5, A6, A9, A10, ORDERS_1:4; ::_thesis: verum end; A15: now__::_thesis:_for_x,_y,_z_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_z_in_Bags_n_&_[x,y]_in_R_&_[y,z]_in_R_holds_ [x,z]_in_R let x, y, z be set ; ::_thesis: ( x in Bags n & y in Bags n & z in Bags n & [x,y] in R & [y,z] in R implies [b1,b3] in R ) assume that A16: x in Bags n and A17: y in Bags n and A18: z in Bags n and A19: [x,y] in R and A20: [y,z] in R ; ::_thesis: [b1,b3] in R consider x9, y9 being bag of n such that A21: x9 = x and A22: y9 = y and A23: ( TotDegree x9 < TotDegree y9 or ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ) by A1, A19; consider y99, z9 being bag of n such that A24: y99 = y and A25: z9 = z and A26: ( TotDegree y99 < TotDegree z9 or ( TotDegree y99 = TotDegree z9 & [y99,z9] in o ) ) by A1, A20; percases ( TotDegree x9 < TotDegree y9 or ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ) by A23; supposeA27: TotDegree x9 < TotDegree y9 ; ::_thesis: [b1,b3] in R now__::_thesis:_[x,z]_in_R percases ( TotDegree y9 < TotDegree z9 or ( TotDegree y9 = TotDegree z9 & [y9,z9] in o ) ) by A22, A24, A26; suppose TotDegree y9 < TotDegree z9 ; ::_thesis: [x,z] in R then TotDegree x9 < TotDegree z9 by A27, XXREAL_0:2; hence [x,z] in R by A1, A16, A18, A21, A25; ::_thesis: verum end; suppose ( TotDegree y9 = TotDegree z9 & [y9,z9] in o ) ; ::_thesis: [x,z] in R hence [x,z] in R by A1, A16, A18, A21, A25, A27; ::_thesis: verum end; end; end; hence [x,z] in R ; ::_thesis: verum end; supposeA28: ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ; ::_thesis: [b1,b3] in R now__::_thesis:_[x,z]_in_R percases ( TotDegree y9 < TotDegree z9 or ( TotDegree y9 = TotDegree z9 & [y9,z9] in o ) ) by A22, A24, A26; suppose TotDegree y9 < TotDegree z9 ; ::_thesis: [x,z] in R hence [x,z] in R by A1, A16, A18, A21, A25, A28; ::_thesis: verum end; suppose ( TotDegree y9 = TotDegree z9 & [y9,z9] in o ) ; ::_thesis: [x,z] in R then [x9,z9] in o by A16, A17, A18, A21, A22, A25, A28, ORDERS_1:5; hence [x,z] in R by A1, A16, A18, A21, A22, A24, A25, A26, A28; ::_thesis: verum end; end; end; hence [x,z] in R ; ::_thesis: verum end; end; end; A29: R is_reflexive_in Bags n by A2, RELAT_2:def_1; A30: R is_antisymmetric_in Bags n by A4, RELAT_2:def_4; A31: R is_transitive_in Bags n by A15, RELAT_2:def_8; A32: dom R = Bags n by A29, ORDERS_1:13; field R = Bags n by A29, ORDERS_1:13; then reconsider R = R as TermOrder of n by A29, A30, A31, A32, PARTFUN1:def_2, RELAT_2:def_9, RELAT_2:def_12, RELAT_2:def_16; take R ; ::_thesis: for a, b being bag of n holds ( [a,b] in R iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) let a, b be bag of n; ::_thesis: ( [a,b] in R iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) hereby ::_thesis: ( ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) implies [a,b] in R ) assume [a,b] in R ; ::_thesis: ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) then ex x9, y9 being bag of n st ( x9 = a & y9 = b & ( TotDegree x9 < TotDegree y9 or ( TotDegree x9 = TotDegree y9 & [x9,y9] in o ) ) ) by A1; hence ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ; ::_thesis: verum end; assume A33: ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ; ::_thesis: [a,b] in R A34: a in Bags n by PRE_POLY:def_12; b in Bags n by PRE_POLY:def_12; hence [a,b] in R by A1, A33, A34; ::_thesis: verum end; uniqueness for b1, b2 being TermOrder of n st ( for a, b being bag of n holds ( [a,b] in b1 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) ) & ( for a, b being bag of n holds ( [a,b] in b2 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) ) holds b1 = b2 proof let IT1, IT2 be TermOrder of n; ::_thesis: ( ( for a, b being bag of n holds ( [a,b] in IT1 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) ) & ( for a, b being bag of n holds ( [a,b] in IT2 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) ) implies IT1 = IT2 ) assume that A35: for a, b being bag of n holds ( [a,b] in IT1 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) and A36: for a, b being bag of n holds ( [a,b] in IT2 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_IT1_implies_[a,b]_in_IT2_)_&_(_[a,b]_in_IT2_implies_[a,b]_in_IT1_)_) let a, b be set ; ::_thesis: ( ( [a,b] in IT1 implies [a,b] in IT2 ) & ( [a,b] in IT2 implies [a,b] in IT1 ) ) hereby ::_thesis: ( [a,b] in IT2 implies [a,b] in IT1 ) assume A37: [a,b] in IT1 ; ::_thesis: [a,b] in IT2 then A38: a in dom IT1 by XTUPLE_0:def_12; b in rng IT1 by A37, XTUPLE_0:def_13; then reconsider a9 = a, b9 = b as bag of n by A38; ( TotDegree a9 < TotDegree b9 or ( TotDegree a9 = TotDegree b9 & [a9,b9] in o ) ) by A35, A37; hence [a,b] in IT2 by A36; ::_thesis: verum end; assume A39: [a,b] in IT2 ; ::_thesis: [a,b] in IT1 then A40: a in dom IT2 by XTUPLE_0:def_12; b in rng IT2 by A39, XTUPLE_0:def_13; then reconsider a9 = a, b9 = b as bag of n by A40; ( TotDegree a9 < TotDegree b9 or ( TotDegree a9 = TotDegree b9 & [a9,b9] in o ) ) by A36, A39; hence [a,b] in IT1 by A35; ::_thesis: verum end; hence IT1 = IT2 by RELAT_1:def_2; ::_thesis: verum end; end; :: deftheorem Def7 defines Graded BAGORDER:def_7_:_ for n being Ordinal for o being TermOrder of n st ( for a, b, c being bag of n st [a,b] in o holds [(a + c),(b + c)] in o ) holds for b3 being TermOrder of n holds ( b3 = Graded o iff for a, b being bag of n holds ( [a,b] in b3 iff ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) ) ); theorem Th26: :: BAGORDER:26 for n being Ordinal for o being TermOrder of n st ( for a, b, c being bag of n st [a,b] in o holds [(a + c),(b + c)] in o ) & o is_strongly_connected_in Bags n holds Graded o is admissible proof let n be Ordinal; ::_thesis: for o being TermOrder of n st ( for a, b, c being bag of n st [a,b] in o holds [(a + c),(b + c)] in o ) & o is_strongly_connected_in Bags n holds Graded o is admissible let o be TermOrder of n; ::_thesis: ( ( for a, b, c being bag of n st [a,b] in o holds [(a + c),(b + c)] in o ) & o is_strongly_connected_in Bags n implies Graded o is admissible ) assume that A1: for a, b, c being bag of n st [a,b] in o holds [(a + c),(b + c)] in o and A2: o is_strongly_connected_in Bags n ; ::_thesis: Graded o is admissible now__::_thesis:_for_x,_y_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_not_[x,y]_in_Graded_o_holds_ [y,x]_in_Graded_o let x, y be set ; ::_thesis: ( x in Bags n & y in Bags n & not [x,y] in Graded o implies [b2,b1] in Graded o ) assume that A3: x in Bags n and A4: y in Bags n ; ::_thesis: ( not [x,y] in Graded o implies [b2,b1] in Graded o ) reconsider x9 = x, y9 = y as bag of n by A3, A4; assume A5: not [x,y] in Graded o ; ::_thesis: [b2,b1] in Graded o then A6: TotDegree x9 >= TotDegree y9 by A1, Def7; percases ( TotDegree y9 < TotDegree x9 or TotDegree y9 = TotDegree x9 ) by A6, XXREAL_0:1; suppose TotDegree y9 < TotDegree x9 ; ::_thesis: [b2,b1] in Graded o hence [y,x] in Graded o by A1, Def7; ::_thesis: verum end; supposeA7: TotDegree y9 = TotDegree x9 ; ::_thesis: [b2,b1] in Graded o then not [x,y] in o by A1, A5, Def7; then [y,x] in o by A2, A3, A4, RELAT_2:def_7; hence [y,x] in Graded o by A1, A7, Def7; ::_thesis: verum end; end; end; hence Graded o is_strongly_connected_in Bags n by RELAT_2:def_7; :: according to BAGORDER:def_5 ::_thesis: ( ( for a being bag of n holds [(EmptyBag n),a] in Graded o ) & ( for a, b, c being bag of n st [a,b] in Graded o holds [(a + c),(b + c)] in Graded o ) ) now__::_thesis:_for_a_being_bag_of_n_holds_[(EmptyBag_n),a]_in_Graded_o let a be bag of n; ::_thesis: [(EmptyBag n),b1] in Graded o A8: TotDegree (EmptyBag n) = 0 by Th15; percases ( a = EmptyBag n or a <> EmptyBag n ) ; suppose a = EmptyBag n ; ::_thesis: [(EmptyBag n),b1] in Graded o hence [(EmptyBag n),a] in Graded o by ORDERS_1:3; ::_thesis: verum end; suppose a <> EmptyBag n ; ::_thesis: [(EmptyBag n),b1] in Graded o then TotDegree a <> 0 by Th15; hence [(EmptyBag n),a] in Graded o by A1, A8, Def7; ::_thesis: verum end; end; end; hence for a being bag of n holds [(EmptyBag n),a] in Graded o ; ::_thesis: for a, b, c being bag of n st [a,b] in Graded o holds [(a + c),(b + c)] in Graded o now__::_thesis:_for_a,_b,_c_being_bag_of_n_st_[a,b]_in_Graded_o_holds_ [(a_+_c),(b_+_c)]_in_Graded_o let a, b, c be bag of n; ::_thesis: ( [a,b] in Graded o implies [(b1 + b3),(b2 + b3)] in Graded o ) assume A9: [a,b] in Graded o ; ::_thesis: [(b1 + b3),(b2 + b3)] in Graded o percases ( TotDegree a < TotDegree b or ( TotDegree a = TotDegree b & [a,b] in o ) ) by A1, A9, Def7; supposeA10: TotDegree a < TotDegree b ; ::_thesis: [(b1 + b3),(b2 + b3)] in Graded o A11: TotDegree (a + c) = (TotDegree a) + (TotDegree c) by Th13; TotDegree (b + c) = (TotDegree b) + (TotDegree c) by Th13; then TotDegree (a + c) < TotDegree (b + c) by A10, A11, XREAL_1:8; hence [(a + c),(b + c)] in Graded o by A1, Def7; ::_thesis: verum end; supposeA12: ( TotDegree a = TotDegree b & [a,b] in o ) ; ::_thesis: [(b1 + b3),(b2 + b3)] in Graded o then TotDegree (a + c) = (TotDegree b) + (TotDegree c) by Th13; then A13: TotDegree (a + c) = TotDegree (b + c) by Th13; [(a + c),(b + c)] in o by A1, A12; hence [(a + c),(b + c)] in Graded o by A1, A13, Def7; ::_thesis: verum end; end; end; hence for a, b, c being bag of n st [a,b] in Graded o holds [(a + c),(b + c)] in Graded o ; ::_thesis: verum end; definition let n be Ordinal; func GrLexOrder n -> TermOrder of n equals :: BAGORDER:def 8 Graded (LexOrder n); coherence Graded (LexOrder n) is TermOrder of n ; func GrInvLexOrder n -> TermOrder of n equals :: BAGORDER:def 9 Graded (InvLexOrder n); coherence Graded (InvLexOrder n) is TermOrder of n ; end; :: deftheorem defines GrLexOrder BAGORDER:def_8_:_ for n being Ordinal holds GrLexOrder n = Graded (LexOrder n); :: deftheorem defines GrInvLexOrder BAGORDER:def_9_:_ for n being Ordinal holds GrInvLexOrder n = Graded (InvLexOrder n); theorem Th27: :: BAGORDER:27 for n being Ordinal holds GrLexOrder n is admissible proof let n be Ordinal; ::_thesis: GrLexOrder n is admissible A1: for a, b, c being bag of n st [a,b] in LexOrder n holds [(a + c),(b + c)] in LexOrder n by Def5; LexOrder n is_strongly_connected_in Bags n by Def5; hence GrLexOrder n is admissible by A1, Th26; ::_thesis: verum end; registration let n be Ordinal; cluster GrLexOrder n -> admissible ; coherence GrLexOrder n is admissible by Th27; end; theorem :: BAGORDER:28 for o being infinite Ordinal holds not GrLexOrder o is well-ordering proof let o be infinite Ordinal; ::_thesis: not GrLexOrder o is well-ordering set R = GrLexOrder o; set r = RelStr(# (Bags o),(GrLexOrder o) #); set ir = the InternalRel of RelStr(# (Bags o),(GrLexOrder o) #); set cr = the carrier of RelStr(# (Bags o),(GrLexOrder o) #); assume GrLexOrder o is well-ordering ; ::_thesis: contradiction then A1: GrLexOrder o is well_founded ; the carrier of RelStr(# (Bags o),(GrLexOrder o) #) = field the InternalRel of RelStr(# (Bags o),(GrLexOrder o) #) by ORDERS_1:15; then the InternalRel of RelStr(# (Bags o),(GrLexOrder o) #) is_well_founded_in the carrier of RelStr(# (Bags o),(GrLexOrder o) #) by A1, WELLORD1:3; then A2: RelStr(# (Bags o),(GrLexOrder o) #) is well_founded by WELLFND1:def_2; defpred S1[ Nat, set ] means $2 = (o --> 0) +* ($1,1); A3: now__::_thesis:_for_n_being_Element_of_NAT_ex_y_being_Element_of_the_carrier_of_RelStr(#_(Bags_o),(GrLexOrder_o)_#)_st_S1[n,y] let n be Element of NAT ; ::_thesis: ex y being Element of the carrier of RelStr(# (Bags o),(GrLexOrder o) #) st S1[n,y] set y = (o --> 0) +* (n,1); A4: dom (o --> 0) = o by FUNCOP_1:13; reconsider y = (o --> 0) +* (n,1) as ManySortedSet of o ; A5: n in omega ; A6: omega c= o by CARD_3:85; now__::_thesis:_for_x_being_set_holds_ (_(_x_in_{n}_implies_y_._x_<>_0_)_&_(_y_._x_<>_0_implies_x_in_{n}_)_) let x be set ; ::_thesis: ( ( x in {n} implies y . x <> 0 ) & ( y . x <> 0 implies x in {n} ) ) hereby ::_thesis: ( y . x <> 0 implies x in {n} ) assume x in {n} ; ::_thesis: y . x <> 0 then x = n by TARSKI:def_1; hence y . x <> 0 by A4, A5, A6, FUNCT_7:31; ::_thesis: verum end; assume that A7: y . x <> 0 and A8: not x in {n} ; ::_thesis: contradiction x <> n by A8, TARSKI:def_1; then A9: y . x = (o --> 0) . x by FUNCT_7:32; percases ( x in dom (o --> 0) or not x in dom (o --> 0) ) ; suppose x in dom (o --> 0) ; ::_thesis: contradiction hence contradiction by A7, A9, FUNCOP_1:7; ::_thesis: verum end; suppose not x in dom (o --> 0) ; ::_thesis: contradiction hence contradiction by A7, A9, FUNCT_1:def_2; ::_thesis: verum end; end; end; then support y = {n} by PRE_POLY:def_7; then y is finite-support by PRE_POLY:def_8; then reconsider y = y as Element of the carrier of RelStr(# (Bags o),(GrLexOrder o) #) by PRE_POLY:def_12; take y = y; ::_thesis: S1[n,y] thus S1[n,y] ; ::_thesis: verum end; consider f being Function of NAT, the carrier of RelStr(# (Bags o),(GrLexOrder o) #) such that A10: for n being Element of NAT holds S1[n,f . n] from FUNCT_2:sch_3(A3); reconsider f = f as sequence of RelStr(# (Bags o),(GrLexOrder o) #) ; f is descending proof let n be Nat; :: according to WELLFND1:def_6 ::_thesis: ( not f . (n + 1) = f . n & [(f . (n + 1)),(f . n)] in the InternalRel of RelStr(# (Bags o),(GrLexOrder o) #) ) reconsider n0 = n as Element of NAT by ORDINAL1:def_12; set fn1 = f . (n0 + 1); set fn = f . n0; A11: f . (n0 + 1) = (o --> 0) +* ((n + 1),1) by A10; A12: f . n0 = (o --> 0) +* (n,1) by A10; reconsider fn1 = f . (n0 + 1) as bag of o ; reconsider fn = f . n0 as bag of o ; A13: n0 in omega ; A14: omega c= o by CARD_3:85; n <> n + 1 ; then A15: fn1 . n = (o --> 0) . n by A11, FUNCT_7:32 .= 0 by A13, A14, FUNCOP_1:7 ; A16: dom (o --> 0) = o by FUNCOP_1:13; then A17: fn . n = 1 by A12, A13, A14, FUNCT_7:31; now__::_thesis:_for_l_being_Ordinal_st_l_in_n_holds_ fn1_._l_=_fn_._l let l be Ordinal; ::_thesis: ( l in n implies fn1 . l = fn . l ) assume A18: l in n ; ::_thesis: fn1 . l = fn . l then A19: l <> n ; n < n + 1 by NAT_1:13; then n in { i where i is Element of NAT : i < n0 + 1 } ; then n in n + 1 by AXIOMS:4; then n c= n + 1 by ORDINAL1:def_2; then l in n + 1 by A18; then l <> n + 1 ; hence fn1 . l = (o --> 0) . l by A11, FUNCT_7:32 .= fn . l by A12, A19, FUNCT_7:32 ; ::_thesis: verum end; then A20: fn1 < fn by A15, A17, PRE_POLY:def_9; thus f . (n + 1) <> f . n by A12, A13, A14, A15, A16, FUNCT_7:31; ::_thesis: [(f . (n + 1)),(f . n)] in the InternalRel of RelStr(# (Bags o),(GrLexOrder o) #) fn1 <=' fn by A20, PRE_POLY:def_10; then A21: [(f . (n + 1)),(f . n)] in LexOrder o by PRE_POLY:def_14; consider tn being FinSequence of NAT such that A22: TotDegree fn = Sum tn and A23: tn = fn * (SgmX ((RelIncl o),(support fn))) by Def4; consider tn1 being FinSequence of NAT such that A24: TotDegree fn1 = Sum tn1 and A25: tn1 = fn1 * (SgmX ((RelIncl o),(support fn1))) by Def4; A26: n + 1 in omega ; omega c= o by CARD_3:85; then reconsider nn = n, n1n = n + 1 as Element of o by A13, A26; A27: field (RelIncl o) = o by WELLORD2:def_1; A28: RelIncl o linearly_orders o by A27, ORDERS_1:19, ORDERS_1:37; now__::_thesis:_for_x_being_set_holds_ (_(_x_in_support_fn1_implies_x_in_{n1n}_)_&_(_x_in_{n1n}_implies_x_in_support_fn1_)_) let x be set ; ::_thesis: ( ( x in support fn1 implies x in {n1n} ) & ( x in {n1n} implies x in support fn1 ) ) hereby ::_thesis: ( x in {n1n} implies x in support fn1 ) assume A29: x in support fn1 ; ::_thesis: x in {n1n} now__::_thesis:_not_x_<>_n1n assume x <> n1n ; ::_thesis: contradiction then fn1 . x = (o --> 0) . x by A11, FUNCT_7:32; then fn1 . x = 0 by A29, FUNCOP_1:7; hence contradiction by A29, PRE_POLY:def_7; ::_thesis: verum end; hence x in {n1n} by TARSKI:def_1; ::_thesis: verum end; assume x in {n1n} ; ::_thesis: x in support fn1 then x = n1n by TARSKI:def_1; then fn1 . x = 1 by A11, A16, FUNCT_7:31; hence x in support fn1 by PRE_POLY:def_7; ::_thesis: verum end; then support fn1 = {n1n} by TARSKI:1; then A30: SgmX ((RelIncl o),(support fn1)) = <*n1n*> by A28, Th11, ORDERS_1:38; A31: dom fn = o by A12, A16, FUNCT_7:30; A32: dom fn1 = o by A11, A16, FUNCT_7:30; now__::_thesis:_for_x_being_set_holds_ (_(_x_in_support_fn_implies_x_in_{nn}_)_&_(_x_in_{nn}_implies_x_in_support_fn_)_) let x be set ; ::_thesis: ( ( x in support fn implies x in {nn} ) & ( x in {nn} implies x in support fn ) ) hereby ::_thesis: ( x in {nn} implies x in support fn ) assume A33: x in support fn ; ::_thesis: x in {nn} now__::_thesis:_not_x_<>_nn assume x <> nn ; ::_thesis: contradiction then fn . x = (o --> 0) . x by A12, FUNCT_7:32; then fn . x = 0 by A33, FUNCOP_1:7; hence contradiction by A33, PRE_POLY:def_7; ::_thesis: verum end; hence x in {nn} by TARSKI:def_1; ::_thesis: verum end; assume x in {nn} ; ::_thesis: x in support fn then x = nn by TARSKI:def_1; then fn . x = 1 by A12, A16, FUNCT_7:31; hence x in support fn by PRE_POLY:def_7; ::_thesis: verum end; then support fn = {nn} by TARSKI:1; then SgmX ((RelIncl o),(support fn)) = <*nn*> by A28, Th11, ORDERS_1:38; then A34: tn = <*(fn . n)*> by A23, A31, FINSEQ_2:34 .= <*1*> by A12, A13, A14, A16, FUNCT_7:31 .= <*(fn1 . n1n)*> by A11, A16, FUNCT_7:31 .= tn1 by A25, A30, A32, FINSEQ_2:34 ; for a, b, c being bag of o st [a,b] in LexOrder o holds [(a + c),(b + c)] in LexOrder o by Def5; hence [(f . (n + 1)),(f . n)] in the InternalRel of RelStr(# (Bags o),(GrLexOrder o) #) by A21, A22, A24, A34, Def7; ::_thesis: verum end; hence contradiction by A2, WELLFND1:14; ::_thesis: verum end; theorem Th29: :: BAGORDER:29 for n being Ordinal holds GrInvLexOrder n is admissible proof let n be Ordinal; ::_thesis: GrInvLexOrder n is admissible A1: for a, b, c being bag of n st [a,b] in InvLexOrder n holds [(a + c),(b + c)] in InvLexOrder n by Def5; InvLexOrder n is_strongly_connected_in Bags n by Def5; hence GrInvLexOrder n is admissible by A1, Th26; ::_thesis: verum end; registration let n be Ordinal; cluster GrInvLexOrder n -> admissible ; coherence GrInvLexOrder n is admissible by Th29; end; theorem :: BAGORDER:30 for o being Ordinal holds GrInvLexOrder o is well-ordering proof let o be Ordinal; ::_thesis: GrInvLexOrder o is well-ordering set gilo = GrInvLexOrder o; set ilo = InvLexOrder o; A1: GrInvLexOrder o is_strongly_connected_in Bags o by Def5; A2: field (GrInvLexOrder o) = Bags o by ORDERS_1:12; then A3: GrInvLexOrder o is_reflexive_in Bags o by RELAT_2:def_9; A4: GrInvLexOrder o is_transitive_in Bags o by A2, RELAT_2:def_16; A5: GrInvLexOrder o is_antisymmetric_in Bags o by A2, RELAT_2:def_12; A6: GrInvLexOrder o is_connected_in field (GrInvLexOrder o) by A1, A2, ORDERS_1:7; A7: for a, b, c being bag of o st [a,b] in InvLexOrder o holds [(a + c),(b + c)] in InvLexOrder o by Def5; now__::_thesis:_for_Y_being_set_st_Y_c=_field_(GrInvLexOrder_o)_&_Y_<>_{}_holds_ ex_a_being_set_st_ (_a_in_Y_&_(GrInvLexOrder_o)_-Seg_a_misses_Y_) let Y be set ; ::_thesis: ( Y c= field (GrInvLexOrder o) & Y <> {} implies ex a being set st ( a in Y & (GrInvLexOrder o) -Seg a misses Y ) ) assume that A8: Y c= field (GrInvLexOrder o) and A9: Y <> {} ; ::_thesis: ex a being set st ( a in Y & (GrInvLexOrder o) -Seg a misses Y ) set TD = { (TotDegree y) where y is Element of Bags o : y in Y } ; consider x being set such that A10: x in Y by A9, XBOOLE_0:def_1; reconsider x = x as Element of Bags o by A8, A10, ORDERS_1:12; A11: TotDegree x in { (TotDegree y) where y is Element of Bags o : y in Y } by A10; { (TotDegree y) where y is Element of Bags o : y in Y } c= NAT proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { (TotDegree y) where y is Element of Bags o : y in Y } or x in NAT ) assume x in { (TotDegree y) where y is Element of Bags o : y in Y } ; ::_thesis: x in NAT then ex y being Element of Bags o st ( x = TotDegree y & y in Y ) ; hence x in NAT by ORDINAL1:def_12; ::_thesis: verum end; then reconsider TD = { (TotDegree y) where y is Element of Bags o : y in Y } as non empty Subset of NAT by A11; set mTD = { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } ; A12: { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } c= field (GrInvLexOrder o) proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } or x in field (GrInvLexOrder o) ) assume x in { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } ; ::_thesis: x in field (GrInvLexOrder o) then ex y being Element of Bags o st ( x = y & y in Y & TotDegree y = min TD ) ; hence x in field (GrInvLexOrder o) by A2; ::_thesis: verum end; min TD in TD by XXREAL_2:def_7; then consider y being Element of Bags o such that A13: TotDegree y = min TD and A14: y in Y ; A15: y in { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } by A13, A14; A16: field (InvLexOrder o) = Bags o by ORDERS_1:12; InvLexOrder o is well-ordering by Th25; then InvLexOrder o well_orders field (InvLexOrder o) by WELLORD1:4; then InvLexOrder o is_well_founded_in field (InvLexOrder o) by WELLORD1:def_5; then consider a being set such that A17: a in { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } and A18: (InvLexOrder o) -Seg a misses { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } by A2, A12, A15, A16, WELLORD1:def_3; A19: ((InvLexOrder o) -Seg a) /\ { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } = {} by A18, XBOOLE_0:def_7; A20: ex a9 being Element of Bags o st ( a9 = a & a9 in Y & TotDegree a9 = min TD ) by A17; take a = a; ::_thesis: ( a in Y & (GrInvLexOrder o) -Seg a misses Y ) thus a in Y by A20; ::_thesis: (GrInvLexOrder o) -Seg a misses Y now__::_thesis:_not_((GrInvLexOrder_o)_-Seg_a)_/\_Y_<>_{} assume ((GrInvLexOrder o) -Seg a) /\ Y <> {} ; ::_thesis: contradiction then consider z being set such that A21: z in ((GrInvLexOrder o) -Seg a) /\ Y by XBOOLE_0:def_1; A22: z in (GrInvLexOrder o) -Seg a by A21, XBOOLE_0:def_4; A23: z in Y by A21, XBOOLE_0:def_4; A24: z <> a by A22, WELLORD1:1; A25: [z,a] in GrInvLexOrder o by A22, WELLORD1:1; reconsider a = a, z = z as Element of Bags o by A8, A20, A23, ORDERS_1:12; percases ( TotDegree z < TotDegree a or TotDegree z = TotDegree a or TotDegree z > TotDegree a ) by XXREAL_0:1; supposeA26: TotDegree z < TotDegree a ; ::_thesis: contradiction TotDegree z in TD by A23; hence contradiction by A20, A26, XXREAL_2:def_7; ::_thesis: verum end; supposeA27: TotDegree z = TotDegree a ; ::_thesis: contradiction then [z,a] in InvLexOrder o by A7, A25, Def7; then A28: z in (InvLexOrder o) -Seg a by A24, WELLORD1:1; z in { y where y is Element of Bags o : ( y in Y & TotDegree y = min TD ) } by A20, A23, A27; hence contradiction by A19, A28, XBOOLE_0:def_4; ::_thesis: verum end; suppose TotDegree z > TotDegree a ; ::_thesis: contradiction hence contradiction by A7, A25, Def7; ::_thesis: verum end; end; end; hence (GrInvLexOrder o) -Seg a misses Y by XBOOLE_0:def_7; ::_thesis: verum end; then GrInvLexOrder o is_well_founded_in field (GrInvLexOrder o) by WELLORD1:def_3; then GrInvLexOrder o well_orders field (GrInvLexOrder o) by A2, A3, A4, A5, A6, WELLORD1:def_5; hence GrInvLexOrder o is well-ordering by WELLORD1:4; ::_thesis: verum end; definition let i, n be Nat; let o1 be TermOrder of (i + 1); let o2 be TermOrder of (n -' (i + 1)); func BlockOrder (i,n,o1,o2) -> TermOrder of n means :Def10: :: BAGORDER:def 10 for p, q being bag of n holds ( [p,q] in it iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ); existence ex b1 being TermOrder of n st for p, q being bag of n holds ( [p,q] in b1 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) proof A1: i + 1 = (i + 1) -' 0 by NAT_D:40; defpred S1[ set , set ] means ex p, q being bag of n st ( $1 = p & $2 = q & ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ); consider BO being Relation of (Bags n),(Bags n) such that A2: for x, y being set holds ( [x,y] in BO iff ( x in Bags n & y in Bags n & S1[x,y] ) ) from RELSET_1:sch_1(); now__::_thesis:_for_x_being_set_st_x_in_Bags_n_holds_ [x,x]_in_BO let x be set ; ::_thesis: ( x in Bags n implies [x,x] in BO ) assume A3: x in Bags n ; ::_thesis: [x,x] in BO reconsider x9 = x as bag of n by A3; A4: (0,(i + 1)) -cut x9 = (0,(i + 1)) -cut x9 ; ((i + 1),n) -cut x9 in Bags (n -' (i + 1)) by PRE_POLY:def_12; then [(((i + 1),n) -cut x9),(((i + 1),n) -cut x9)] in o2 by ORDERS_1:3; hence [x,x] in BO by A2, A3, A4; ::_thesis: verum end; then A5: BO is_reflexive_in Bags n by RELAT_2:def_1; now__::_thesis:_for_x,_y_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_[x,y]_in_BO_&_[y,x]_in_BO_holds_ x_=_y let x, y be set ; ::_thesis: ( x in Bags n & y in Bags n & [x,y] in BO & [y,x] in BO implies b1 = b2 ) assume that x in Bags n and y in Bags n and A6: [x,y] in BO and A7: [y,x] in BO ; ::_thesis: b1 = b2 consider p, q being bag of n such that A8: x = p and A9: y = q and A10: ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) by A2, A6; A11: ex q9, p9 being bag of n st ( y = q9 & x = p9 & ( ( (0,(i + 1)) -cut q9 <> (0,(i + 1)) -cut p9 & [((0,(i + 1)) -cut q9),((0,(i + 1)) -cut p9)] in o1 ) or ( (0,(i + 1)) -cut q9 = (0,(i + 1)) -cut p9 & [(((i + 1),n) -cut q9),(((i + 1),n) -cut p9)] in o2 ) ) ) by A2, A7; set CUTP1 = (0,(i + 1)) -cut p; set CUTP2 = ((i + 1),n) -cut p; set CUTQ1 = (0,(i + 1)) -cut q; set CUTQ2 = ((i + 1),n) -cut q; A12: (0,(i + 1)) -cut p in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A13: (0,(i + 1)) -cut q in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A14: ((i + 1),n) -cut p in Bags (n -' (i + 1)) by PRE_POLY:def_12; A15: ((i + 1),n) -cut q in Bags (n -' (i + 1)) by PRE_POLY:def_12; percases ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) by A10; supposeA16: ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) ; ::_thesis: b1 = b2 now__::_thesis:_x_=_y percases ( ( (0,(i + 1)) -cut q <> (0,(i + 1)) -cut p & [((0,(i + 1)) -cut q),((0,(i + 1)) -cut p)] in o1 ) or ( (0,(i + 1)) -cut q = (0,(i + 1)) -cut p & [(((i + 1),n) -cut q),(((i + 1),n) -cut p)] in o2 ) ) by A8, A9, A11; suppose ( (0,(i + 1)) -cut q <> (0,(i + 1)) -cut p & [((0,(i + 1)) -cut q),((0,(i + 1)) -cut p)] in o1 ) ; ::_thesis: x = y hence x = y by A1, A12, A13, A16, ORDERS_1:4; ::_thesis: verum end; suppose ( (0,(i + 1)) -cut q = (0,(i + 1)) -cut p & [(((i + 1),n) -cut q),(((i + 1),n) -cut p)] in o2 ) ; ::_thesis: x = y hence x = y by A16; ::_thesis: verum end; end; end; hence x = y ; ::_thesis: verum end; supposeA17: ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ; ::_thesis: b1 = b2 now__::_thesis:_x_=_y percases ( ( (0,(i + 1)) -cut q <> (0,(i + 1)) -cut p & [((0,(i + 1)) -cut q),((0,(i + 1)) -cut p)] in o1 ) or ( (0,(i + 1)) -cut q = (0,(i + 1)) -cut p & [(((i + 1),n) -cut q),(((i + 1),n) -cut p)] in o2 ) ) by A8, A9, A11; suppose ( (0,(i + 1)) -cut q <> (0,(i + 1)) -cut p & [((0,(i + 1)) -cut q),((0,(i + 1)) -cut p)] in o1 ) ; ::_thesis: x = y hence x = y by A17; ::_thesis: verum end; suppose ( (0,(i + 1)) -cut q = (0,(i + 1)) -cut p & [(((i + 1),n) -cut q),(((i + 1),n) -cut p)] in o2 ) ; ::_thesis: x = y then ((i + 1),n) -cut q = ((i + 1),n) -cut p by A14, A15, A17, ORDERS_1:4; hence x = y by A8, A9, A17, Th5; ::_thesis: verum end; end; end; hence x = y ; ::_thesis: verum end; end; end; then A18: BO is_antisymmetric_in Bags n by RELAT_2:def_4; now__::_thesis:_for_x,_y,_z_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_z_in_Bags_n_&_[x,y]_in_BO_&_[y,z]_in_BO_holds_ [x,z]_in_BO let x, y, z be set ; ::_thesis: ( x in Bags n & y in Bags n & z in Bags n & [x,y] in BO & [y,z] in BO implies [b1,b3] in BO ) assume that A19: x in Bags n and y in Bags n and A20: z in Bags n and A21: [x,y] in BO and A22: [y,z] in BO ; ::_thesis: [b1,b3] in BO consider x9, y9 being bag of n such that A23: x9 = x and A24: y9 = y and A25: ( ( (0,(i + 1)) -cut x9 <> (0,(i + 1)) -cut y9 & [((0,(i + 1)) -cut x9),((0,(i + 1)) -cut y9)] in o1 ) or ( (0,(i + 1)) -cut x9 = (0,(i + 1)) -cut y9 & [(((i + 1),n) -cut x9),(((i + 1),n) -cut y9)] in o2 ) ) by A2, A21; consider y99, z9 being bag of n such that A26: y99 = y and A27: z9 = z and A28: ( ( (0,(i + 1)) -cut y99 <> (0,(i + 1)) -cut z9 & [((0,(i + 1)) -cut y99),((0,(i + 1)) -cut z9)] in o1 ) or ( (0,(i + 1)) -cut y99 = (0,(i + 1)) -cut z9 & [(((i + 1),n) -cut y99),(((i + 1),n) -cut z9)] in o2 ) ) by A2, A22; set CUTX1 = (0,(i + 1)) -cut x9; set CUTX2 = ((i + 1),n) -cut x9; set CUTY1 = (0,(i + 1)) -cut y9; set CUTY2 = ((i + 1),n) -cut y9; set CUTZ1 = (0,(i + 1)) -cut z9; set CUTZ2 = ((i + 1),n) -cut z9; A29: (0,(i + 1)) -cut x9 in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A30: (0,(i + 1)) -cut y9 in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A31: (0,(i + 1)) -cut z9 in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A32: ((i + 1),n) -cut x9 in Bags (n -' (i + 1)) by PRE_POLY:def_12; A33: ((i + 1),n) -cut y9 in Bags (n -' (i + 1)) by PRE_POLY:def_12; A34: ((i + 1),n) -cut z9 in Bags (n -' (i + 1)) by PRE_POLY:def_12; percases ( ( (0,(i + 1)) -cut x9 <> (0,(i + 1)) -cut y9 & [((0,(i + 1)) -cut x9),((0,(i + 1)) -cut y9)] in o1 ) or ( (0,(i + 1)) -cut x9 = (0,(i + 1)) -cut y9 & [(((i + 1),n) -cut x9),(((i + 1),n) -cut y9)] in o2 ) ) by A25; supposeA35: ( (0,(i + 1)) -cut x9 <> (0,(i + 1)) -cut y9 & [((0,(i + 1)) -cut x9),((0,(i + 1)) -cut y9)] in o1 ) ; ::_thesis: [b1,b3] in BO now__::_thesis:_[x,z]_in_BO percases ( ( (0,(i + 1)) -cut y9 <> (0,(i + 1)) -cut z9 & [((0,(i + 1)) -cut y9),((0,(i + 1)) -cut z9)] in o1 ) or ( (0,(i + 1)) -cut y9 = (0,(i + 1)) -cut z9 & [(((i + 1),n) -cut y9),(((i + 1),n) -cut z9)] in o2 ) ) by A24, A26, A28; supposeA36: ( (0,(i + 1)) -cut y9 <> (0,(i + 1)) -cut z9 & [((0,(i + 1)) -cut y9),((0,(i + 1)) -cut z9)] in o1 ) ; ::_thesis: [x,z] in BO then A37: [((0,(i + 1)) -cut x9),((0,(i + 1)) -cut z9)] in o1 by A1, A29, A30, A31, A35, ORDERS_1:5; now__::_thesis:_[x,z]_in_BO percases ( (0,(i + 1)) -cut x9 <> (0,(i + 1)) -cut z9 or (0,(i + 1)) -cut x9 = (0,(i + 1)) -cut z9 ) ; suppose (0,(i + 1)) -cut x9 <> (0,(i + 1)) -cut z9 ; ::_thesis: [x,z] in BO hence [x,z] in BO by A2, A19, A20, A23, A27, A37; ::_thesis: verum end; suppose (0,(i + 1)) -cut x9 = (0,(i + 1)) -cut z9 ; ::_thesis: [x,z] in BO hence [x,z] in BO by A1, A29, A30, A35, A36, ORDERS_1:4; ::_thesis: verum end; end; end; hence [x,z] in BO ; ::_thesis: verum end; suppose ( (0,(i + 1)) -cut y9 = (0,(i + 1)) -cut z9 & [(((i + 1),n) -cut y9),(((i + 1),n) -cut z9)] in o2 ) ; ::_thesis: [x,z] in BO hence [x,z] in BO by A2, A19, A20, A23, A27, A35; ::_thesis: verum end; end; end; hence [x,z] in BO ; ::_thesis: verum end; supposeA38: ( (0,(i + 1)) -cut x9 = (0,(i + 1)) -cut y9 & [(((i + 1),n) -cut x9),(((i + 1),n) -cut y9)] in o2 ) ; ::_thesis: [b1,b3] in BO now__::_thesis:_[x,z]_in_BO percases ( ( (0,(i + 1)) -cut y9 <> (0,(i + 1)) -cut z9 & [((0,(i + 1)) -cut y9),((0,(i + 1)) -cut z9)] in o1 ) or ( (0,(i + 1)) -cut y9 = (0,(i + 1)) -cut z9 & [(((i + 1),n) -cut y9),(((i + 1),n) -cut z9)] in o2 ) ) by A24, A26, A28; suppose ( (0,(i + 1)) -cut y9 <> (0,(i + 1)) -cut z9 & [((0,(i + 1)) -cut y9),((0,(i + 1)) -cut z9)] in o1 ) ; ::_thesis: [x,z] in BO hence [x,z] in BO by A2, A19, A20, A23, A27, A38; ::_thesis: verum end; supposeA39: ( (0,(i + 1)) -cut y9 = (0,(i + 1)) -cut z9 & [(((i + 1),n) -cut y9),(((i + 1),n) -cut z9)] in o2 ) ; ::_thesis: [x,z] in BO then [(((i + 1),n) -cut x9),(((i + 1),n) -cut z9)] in o2 by A32, A33, A34, A38, ORDERS_1:5; hence [x,z] in BO by A2, A19, A20, A23, A27, A38, A39; ::_thesis: verum end; end; end; hence [x,z] in BO ; ::_thesis: verum end; end; end; then A40: BO is_transitive_in Bags n by RELAT_2:def_8; A41: dom BO = Bags n by A5, ORDERS_1:13; field BO = Bags n by A5, ORDERS_1:13; then reconsider BO = BO as TermOrder of n by A5, A18, A40, A41, PARTFUN1:def_2, RELAT_2:def_9, RELAT_2:def_12, RELAT_2:def_16; take BO ; ::_thesis: for p, q being bag of n holds ( [p,q] in BO iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) let p, q be bag of n; ::_thesis: ( [p,q] in BO iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) hereby ::_thesis: ( ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) implies [p,q] in BO ) assume [p,q] in BO ; ::_thesis: ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) then ex p9, q9 being bag of n st ( p9 = p & q9 = q & ( ( (0,(i + 1)) -cut p9 <> (0,(i + 1)) -cut q9 & [((0,(i + 1)) -cut p9),((0,(i + 1)) -cut q9)] in o1 ) or ( (0,(i + 1)) -cut p9 = (0,(i + 1)) -cut q9 & [(((i + 1),n) -cut p9),(((i + 1),n) -cut q9)] in o2 ) ) ) by A2; hence ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ; ::_thesis: verum end; assume A42: ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ; ::_thesis: [p,q] in BO A43: p in Bags n by PRE_POLY:def_12; q in Bags n by PRE_POLY:def_12; hence [p,q] in BO by A2, A42, A43; ::_thesis: verum end; uniqueness for b1, b2 being TermOrder of n st ( for p, q being bag of n holds ( [p,q] in b1 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) ) & ( for p, q being bag of n holds ( [p,q] in b2 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) ) holds b1 = b2 proof let IT1, IT2 be TermOrder of n; ::_thesis: ( ( for p, q being bag of n holds ( [p,q] in IT1 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) ) & ( for p, q being bag of n holds ( [p,q] in IT2 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) ) implies IT1 = IT2 ) assume that A44: for p, q being bag of n holds ( [p,q] in IT1 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) and A45: for p, q being bag of n holds ( [p,q] in IT2 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_IT1_implies_[a,b]_in_IT2_)_&_(_[a,b]_in_IT2_implies_[a,b]_in_IT1_)_) let a, b be set ; ::_thesis: ( ( [a,b] in IT1 implies [a,b] in IT2 ) & ( [a,b] in IT2 implies [a,b] in IT1 ) ) hereby ::_thesis: ( [a,b] in IT2 implies [a,b] in IT1 ) assume A46: [a,b] in IT1 ; ::_thesis: [a,b] in IT2 then A47: a in dom IT1 by XTUPLE_0:def_12; b in rng IT1 by A46, XTUPLE_0:def_13; then reconsider p = a, q = b as bag of n by A47; ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) by A44, A46; hence [a,b] in IT2 by A45; ::_thesis: verum end; assume A48: [a,b] in IT2 ; ::_thesis: [a,b] in IT1 then A49: a in dom IT2 by XTUPLE_0:def_12; b in rng IT2 by A48, XTUPLE_0:def_13; then reconsider p = a, q = b as bag of n by A49; ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) by A45, A48; hence [a,b] in IT1 by A44; ::_thesis: verum end; hence IT1 = IT2 by RELAT_1:def_2; ::_thesis: verum end; end; :: deftheorem Def10 defines BlockOrder BAGORDER:def_10_:_ for i, n being Nat for o1 being TermOrder of (i + 1) for o2 being TermOrder of (n -' (i + 1)) for b5 being TermOrder of n holds ( b5 = BlockOrder (i,n,o1,o2) iff for p, q being bag of n holds ( [p,q] in b5 iff ( ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q & [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) or ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q & [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) ) ) ); theorem :: BAGORDER:31 for i, n being Nat for o1 being TermOrder of (i + 1) for o2 being TermOrder of (n -' (i + 1)) st o1 is admissible & o2 is admissible holds BlockOrder (i,n,o1,o2) is admissible proof let i, n be Nat; ::_thesis: for o1 being TermOrder of (i + 1) for o2 being TermOrder of (n -' (i + 1)) st o1 is admissible & o2 is admissible holds BlockOrder (i,n,o1,o2) is admissible let o1 be TermOrder of (i + 1); ::_thesis: for o2 being TermOrder of (n -' (i + 1)) st o1 is admissible & o2 is admissible holds BlockOrder (i,n,o1,o2) is admissible let o2 be TermOrder of (n -' (i + 1)); ::_thesis: ( o1 is admissible & o2 is admissible implies BlockOrder (i,n,o1,o2) is admissible ) assume that A1: o1 is admissible and A2: o2 is admissible ; ::_thesis: BlockOrder (i,n,o1,o2) is admissible A3: i + 1 = (i + 1) -' 0 by NAT_D:40; then A4: o1 is_strongly_connected_in Bags ((i + 1) -' 0) by A1, Def5; A5: o2 is_strongly_connected_in Bags (n -' (i + 1)) by A2, Def5; set BO = BlockOrder (i,n,o1,o2); now__::_thesis:_(_BlockOrder_(i,n,o1,o2)_is_strongly_connected_in_Bags_n_&_(_for_a_being_bag_of_n_holds_ (_[(EmptyBag_n),a]_in_BlockOrder_(i,n,o1,o2)_&_(_for_a,_b,_c_being_bag_of_n_st_[a,b]_in_BlockOrder_(i,n,o1,o2)_holds_ [(a_+_c),(b_+_c)]_in_BlockOrder_(i,n,o1,o2)_)_)_)_) now__::_thesis:_for_x,_y_being_set_st_x_in_Bags_n_&_y_in_Bags_n_&_not_[x,y]_in_BlockOrder_(i,n,o1,o2)_holds_ [y,x]_in_BlockOrder_(i,n,o1,o2) let x, y be set ; ::_thesis: ( x in Bags n & y in Bags n & not [x,y] in BlockOrder (i,n,o1,o2) implies [b2,b1] in BlockOrder (i,n,o1,o2) ) assume that A6: x in Bags n and A7: y in Bags n ; ::_thesis: ( not [x,y] in BlockOrder (i,n,o1,o2) implies [b2,b1] in BlockOrder (i,n,o1,o2) ) reconsider p = x, q = y as bag of n by A6, A7; set CUTP1 = (0,(i + 1)) -cut p; set CUTP2 = ((i + 1),n) -cut p; set CUTQ1 = (0,(i + 1)) -cut q; set CUTQ2 = ((i + 1),n) -cut q; A8: (0,(i + 1)) -cut p in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A9: (0,(i + 1)) -cut q in Bags ((i + 1) -' 0) by PRE_POLY:def_12; A10: ((i + 1),n) -cut p in Bags (n -' (i + 1)) by PRE_POLY:def_12; A11: ((i + 1),n) -cut q in Bags (n -' (i + 1)) by PRE_POLY:def_12; assume A12: not [x,y] in BlockOrder (i,n,o1,o2) ; ::_thesis: [b2,b1] in BlockOrder (i,n,o1,o2) percases ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q or not [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ) by A12, Def10; supposeA13: (0,(i + 1)) -cut p = (0,(i + 1)) -cut q ; ::_thesis: [b2,b1] in BlockOrder (i,n,o1,o2) now__::_thesis:_[y,x]_in_BlockOrder_(i,n,o1,o2) percases ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q or not [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) by A12, Def10; suppose (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q ; ::_thesis: [y,x] in BlockOrder (i,n,o1,o2) hence [y,x] in BlockOrder (i,n,o1,o2) by A13; ::_thesis: verum end; suppose not [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ; ::_thesis: [y,x] in BlockOrder (i,n,o1,o2) then [(((i + 1),n) -cut q),(((i + 1),n) -cut p)] in o2 by A5, A10, A11, RELAT_2:def_7; hence [y,x] in BlockOrder (i,n,o1,o2) by A13, Def10; ::_thesis: verum end; end; end; hence [y,x] in BlockOrder (i,n,o1,o2) ; ::_thesis: verum end; suppose not [((0,(i + 1)) -cut p),((0,(i + 1)) -cut q)] in o1 ; ::_thesis: [b2,b1] in BlockOrder (i,n,o1,o2) then A14: [((0,(i + 1)) -cut q),((0,(i + 1)) -cut p)] in o1 by A4, A8, A9, RELAT_2:def_7; now__::_thesis:_[y,x]_in_BlockOrder_(i,n,o1,o2) percases ( (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q or not [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ) by A12, Def10; suppose (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q ; ::_thesis: [y,x] in BlockOrder (i,n,o1,o2) hence [y,x] in BlockOrder (i,n,o1,o2) by A14, Def10; ::_thesis: verum end; suppose not [(((i + 1),n) -cut p),(((i + 1),n) -cut q)] in o2 ; ::_thesis: [y,x] in BlockOrder (i,n,o1,o2) then A15: [(((i + 1),n) -cut q),(((i + 1),n) -cut p)] in o2 by A5, A10, A11, RELAT_2:def_7; now__::_thesis:_[y,x]_in_BlockOrder_(i,n,o1,o2) percases ( (0,(i + 1)) -cut p = (0,(i + 1)) -cut q or (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q ) ; suppose (0,(i + 1)) -cut p = (0,(i + 1)) -cut q ; ::_thesis: [y,x] in BlockOrder (i,n,o1,o2) hence [y,x] in BlockOrder (i,n,o1,o2) by A15, Def10; ::_thesis: verum end; suppose (0,(i + 1)) -cut p <> (0,(i + 1)) -cut q ; ::_thesis: [y,x] in BlockOrder (i,n,o1,o2) hence [y,x] in BlockOrder (i,n,o1,o2) by A14, Def10; ::_thesis: verum end; end; end; hence [y,x] in BlockOrder (i,n,o1,o2) ; ::_thesis: verum end; end; end; hence [y,x] in BlockOrder (i,n,o1,o2) ; ::_thesis: verum end; end; end; hence BlockOrder (i,n,o1,o2) is_strongly_connected_in Bags n by RELAT_2:def_7; ::_thesis: for a being bag of n holds ( [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) & ( for a, b, c being bag of n st [a,b] in BlockOrder (i,n,o1,o2) holds [(b6 + b8),(b7 + b8)] in BlockOrder (i,n,o1,o2) ) ) let a be bag of n; ::_thesis: ( [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) & ( for a, b, c being bag of n st [a,b] in BlockOrder (i,n,o1,o2) holds [(b5 + b7),(b6 + b7)] in BlockOrder (i,n,o1,o2) ) ) set CUTE1 = (0,(i + 1)) -cut (EmptyBag n); set CUTA1 = (0,(i + 1)) -cut a; set CUTE2 = ((i + 1),n) -cut (EmptyBag n); set CUTA2 = ((i + 1),n) -cut a; now__::_thesis:_[(EmptyBag_n),a]_in_BlockOrder_(i,n,o1,o2) percases ( (0,(i + 1)) -cut (EmptyBag n) <> (0,(i + 1)) -cut a or (0,(i + 1)) -cut (EmptyBag n) = (0,(i + 1)) -cut a ) ; supposeA16: (0,(i + 1)) -cut (EmptyBag n) <> (0,(i + 1)) -cut a ; ::_thesis: [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) (0,(i + 1)) -cut (EmptyBag n) = EmptyBag ((i + 1) -' 0) by Th16; then [((0,(i + 1)) -cut (EmptyBag n)),((0,(i + 1)) -cut a)] in o1 by A1, A3, Def5; hence [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) by A16, Def10; ::_thesis: verum end; supposeA17: (0,(i + 1)) -cut (EmptyBag n) = (0,(i + 1)) -cut a ; ::_thesis: [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) ((i + 1),n) -cut (EmptyBag n) = EmptyBag (n -' (i + 1)) by Th16; then [(((i + 1),n) -cut (EmptyBag n)),(((i + 1),n) -cut a)] in o2 by A2, Def5; hence [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) by A17, Def10; ::_thesis: verum end; end; end; hence [(EmptyBag n),a] in BlockOrder (i,n,o1,o2) ; ::_thesis: for a, b, c being bag of n st [a,b] in BlockOrder (i,n,o1,o2) holds [(b5 + b7),(b6 + b7)] in BlockOrder (i,n,o1,o2) let a, b, c be bag of n; ::_thesis: ( [a,b] in BlockOrder (i,n,o1,o2) implies [(b2 + b4),(b3 + b4)] in BlockOrder (i,n,o1,o2) ) assume A18: [a,b] in BlockOrder (i,n,o1,o2) ; ::_thesis: [(b2 + b4),(b3 + b4)] in BlockOrder (i,n,o1,o2) set CUTA1 = (0,(i + 1)) -cut a; set CUTA2 = ((i + 1),n) -cut a; set CUTB1 = (0,(i + 1)) -cut b; set CUTB2 = ((i + 1),n) -cut b; set CUTC1 = (0,(i + 1)) -cut c; set CUTC2 = ((i + 1),n) -cut c; set CUTAC1 = (0,(i + 1)) -cut (a + c); set CUTBC1 = (0,(i + 1)) -cut (b + c); set CUTAC2 = ((i + 1),n) -cut (a + c); set CUTBC2 = ((i + 1),n) -cut (b + c); percases ( ( (0,(i + 1)) -cut a <> (0,(i + 1)) -cut b & [((0,(i + 1)) -cut a),((0,(i + 1)) -cut b)] in o1 ) or ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & [(((i + 1),n) -cut a),(((i + 1),n) -cut b)] in o2 ) ) by A18, Def10; supposeA19: ( (0,(i + 1)) -cut a <> (0,(i + 1)) -cut b & [((0,(i + 1)) -cut a),((0,(i + 1)) -cut b)] in o1 ) ; ::_thesis: [(b2 + b4),(b3 + b4)] in BlockOrder (i,n,o1,o2) then [(((0,(i + 1)) -cut a) + ((0,(i + 1)) -cut c)),(((0,(i + 1)) -cut b) + ((0,(i + 1)) -cut c))] in o1 by A1, A3, Def5; then [((0,(i + 1)) -cut (a + c)),(((0,(i + 1)) -cut b) + ((0,(i + 1)) -cut c))] in o1 by Th17; then A20: [((0,(i + 1)) -cut (a + c)),((0,(i + 1)) -cut (b + c))] in o1 by Th17; now__::_thesis:_not_((0,(i_+_1))_-cut_a)_+_((0,(i_+_1))_-cut_c)_=_((0,(i_+_1))_-cut_b)_+_((0,(i_+_1))_-cut_c) assume A21: ((0,(i + 1)) -cut a) + ((0,(i + 1)) -cut c) = ((0,(i + 1)) -cut b) + ((0,(i + 1)) -cut c) ; ::_thesis: contradiction (((0,(i + 1)) -cut a) + ((0,(i + 1)) -cut c)) -' ((0,(i + 1)) -cut c) = (0,(i + 1)) -cut a by PRE_POLY:48; hence contradiction by A19, A21, PRE_POLY:48; ::_thesis: verum end; then (0,(i + 1)) -cut (a + c) <> ((0,(i + 1)) -cut b) + ((0,(i + 1)) -cut c) by Th17; then (0,(i + 1)) -cut (a + c) <> (0,(i + 1)) -cut (b + c) by Th17; hence [(a + c),(b + c)] in BlockOrder (i,n,o1,o2) by A20, Def10; ::_thesis: verum end; supposeA22: ( (0,(i + 1)) -cut a = (0,(i + 1)) -cut b & [(((i + 1),n) -cut a),(((i + 1),n) -cut b)] in o2 ) ; ::_thesis: [(b2 + b4),(b3 + b4)] in BlockOrder (i,n,o1,o2) then [((((i + 1),n) -cut a) + (((i + 1),n) -cut c)),((((i + 1),n) -cut b) + (((i + 1),n) -cut c))] in o2 by A2, Def5; then [(((i + 1),n) -cut (a + c)),((((i + 1),n) -cut b) + (((i + 1),n) -cut c))] in o2 by Th17; then A23: [(((i + 1),n) -cut (a + c)),(((i + 1),n) -cut (b + c))] in o2 by Th17; (0,(i + 1)) -cut (a + c) = ((0,(i + 1)) -cut b) + ((0,(i + 1)) -cut c) by A22, Th17; then (0,(i + 1)) -cut (a + c) = (0,(i + 1)) -cut (b + c) by Th17; hence [(a + c),(b + c)] in BlockOrder (i,n,o1,o2) by A23, Def10; ::_thesis: verum end; end; end; hence BlockOrder (i,n,o1,o2) is admissible by Def5; ::_thesis: verum end; definition let n be Nat; func NaivelyOrderedBags n -> strict RelStr means :Def11: :: BAGORDER:def 11 ( the carrier of it = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of it iff x divides y ) ) ); existence ex b1 being strict RelStr st ( the carrier of b1 = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of b1 iff x divides y ) ) ) proof defpred S1[ set , set ] means ex x9, y9 being bag of n st ( x9 = $1 & y9 = $2 & x9 divides y9 ); consider IR being Relation of (Bags n),(Bags n) such that A1: for x, y being set holds ( [x,y] in IR iff ( x in Bags n & y in Bags n & S1[x,y] ) ) from RELSET_1:sch_1(); set IT = RelStr(# (Bags n),IR #); reconsider IT = RelStr(# (Bags n),IR #) as strict RelStr ; take IT ; ::_thesis: ( the carrier of IT = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of IT iff x divides y ) ) ) thus the carrier of IT = Bags n ; ::_thesis: for x, y being bag of n holds ( [x,y] in the InternalRel of IT iff x divides y ) let x, y be bag of n; ::_thesis: ( [x,y] in the InternalRel of IT iff x divides y ) hereby ::_thesis: ( x divides y implies [x,y] in the InternalRel of IT ) assume [x,y] in the InternalRel of IT ; ::_thesis: x divides y then ex x9, y9 being bag of n st ( x9 = x & y9 = y & x9 divides y9 ) by A1; hence x divides y ; ::_thesis: verum end; assume A2: x divides y ; ::_thesis: [x,y] in the InternalRel of IT A3: x in Bags n by PRE_POLY:def_12; y in Bags n by PRE_POLY:def_12; hence [x,y] in the InternalRel of IT by A1, A2, A3; ::_thesis: verum end; uniqueness for b1, b2 being strict RelStr st the carrier of b1 = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of b1 iff x divides y ) ) & the carrier of b2 = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of b2 iff x divides y ) ) holds b1 = b2 proof let IT1, IT2 be strict RelStr ; ::_thesis: ( the carrier of IT1 = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of IT1 iff x divides y ) ) & the carrier of IT2 = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of IT2 iff x divides y ) ) implies IT1 = IT2 ) assume that A4: the carrier of IT1 = Bags n and A5: for x, y being bag of n holds ( [x,y] in the InternalRel of IT1 iff x divides y ) and A6: the carrier of IT2 = Bags n and A7: for x, y being bag of n holds ( [x,y] in the InternalRel of IT2 iff x divides y ) ; ::_thesis: IT1 = IT2 now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_the_InternalRel_of_IT1_implies_[a,b]_in_the_InternalRel_of_IT2_)_&_(_[a,b]_in_the_InternalRel_of_IT2_implies_[a,b]_in_the_InternalRel_of_IT1_)_) let a, b be set ; ::_thesis: ( ( [a,b] in the InternalRel of IT1 implies [a,b] in the InternalRel of IT2 ) & ( [a,b] in the InternalRel of IT2 implies [a,b] in the InternalRel of IT1 ) ) set z = [a,b]; hereby ::_thesis: ( [a,b] in the InternalRel of IT2 implies [a,b] in the InternalRel of IT1 ) assume A8: [a,b] in the InternalRel of IT1 ; ::_thesis: [a,b] in the InternalRel of IT2 then consider a9, b9 being set such that A9: [a,b] = [a9,b9] and A10: a9 in Bags n and A11: b9 in Bags n by A4, RELSET_1:2; reconsider a9 = a9, b9 = b9 as bag of n by A10, A11; a9 divides b9 by A5, A8, A9; hence [a,b] in the InternalRel of IT2 by A7, A9; ::_thesis: verum end; assume A12: [a,b] in the InternalRel of IT2 ; ::_thesis: [a,b] in the InternalRel of IT1 set z = [a,b]; consider a9, b9 being set such that A13: [a,b] = [a9,b9] and A14: a9 in Bags n and A15: b9 in Bags n by A6, A12, RELSET_1:2; reconsider a9 = a9, b9 = b9 as bag of n by A14, A15; a9 divides b9 by A7, A12, A13; hence [a,b] in the InternalRel of IT1 by A5, A13; ::_thesis: verum end; hence IT1 = IT2 by A4, A6, RELAT_1:def_2; ::_thesis: verum end; end; :: deftheorem Def11 defines NaivelyOrderedBags BAGORDER:def_11_:_ for n being Nat for b2 being strict RelStr holds ( b2 = NaivelyOrderedBags n iff ( the carrier of b2 = Bags n & ( for x, y being bag of n holds ( [x,y] in the InternalRel of b2 iff x divides y ) ) ) ); theorem Th32: :: BAGORDER:32 for n being Nat holds the carrier of (product (n --> OrderedNAT)) = Bags n proof let n be Nat; ::_thesis: the carrier of (product (n --> OrderedNAT)) = Bags n set X = the carrier of (product (n --> OrderedNAT)); A1: the carrier of (product (n --> OrderedNAT)) = product (Carrier (n --> OrderedNAT)) by YELLOW_1:def_4; now__::_thesis:_for_x_being_set_holds_ (_(_x_in_the_carrier_of_(product_(n_-->_OrderedNAT))_implies_x_in_Bags_n_)_&_(_x_in_Bags_n_implies_x_in_the_carrier_of_(product_(n_-->_OrderedNAT))_)_) let x be set ; ::_thesis: ( ( x in the carrier of (product (n --> OrderedNAT)) implies x in Bags n ) & ( x in Bags n implies x in the carrier of (product (n --> OrderedNAT)) ) ) hereby ::_thesis: ( x in Bags n implies x in the carrier of (product (n --> OrderedNAT)) ) assume x in the carrier of (product (n --> OrderedNAT)) ; ::_thesis: x in Bags n then consider g being Function such that A2: x = g and A3: dom g = dom (Carrier (n --> OrderedNAT)) and A4: for i being set st i in dom (Carrier (n --> OrderedNAT)) holds g . i in (Carrier (n --> OrderedNAT)) . i by A1, CARD_3:def_5; A5: dom g = n by A3, PARTFUN1:def_2; A6: rng g c= NAT proof let z be set ; :: according to TARSKI:def_3 ::_thesis: ( not z in rng g or z in NAT ) assume z in rng g ; ::_thesis: z in NAT then consider y being set such that A7: y in dom g and A8: z = g . y by FUNCT_1:def_3; A9: z in (Carrier (n --> OrderedNAT)) . y by A3, A4, A7, A8; ex R being 1-sorted st ( R = (n --> OrderedNAT) . y & (Carrier (n --> OrderedNAT)) . y = the carrier of R ) by A5, A7, PRALG_1:def_13; hence z in NAT by A5, A7, A9, DICKSON:def_15, FUNCOP_1:7; ::_thesis: verum end; A10: dom g = dom (Carrier (n --> OrderedNAT)) by A3 .= n by PARTFUN1:def_2 ; A11: g is natural-valued by A6, VALUED_0:def_6; g is ManySortedSet of n by A10, PARTFUN1:def_2, RELAT_1:def_18; hence x in Bags n by A2, A11, PRE_POLY:def_12; ::_thesis: verum end; assume x in Bags n ; ::_thesis: x in the carrier of (product (n --> OrderedNAT)) then reconsider g = x as natural-valued finite-support ManySortedSet of n ; A12: dom g = n by PARTFUN1:def_2; now__::_thesis:_ex_g_being_natural-valued_finite-support_ManySortedSet_of_n_st_ (_x_=_g_&_dom_g_=_dom_(Carrier_(n_-->_OrderedNAT))_&_(_for_i_being_set_st_i_in_dom_(Carrier_(n_-->_OrderedNAT))_holds_ g_._i_in_(Carrier_(n_-->_OrderedNAT))_._i_)_) take g = g; ::_thesis: ( x = g & dom g = dom (Carrier (n --> OrderedNAT)) & ( for i being set st i in dom (Carrier (n --> OrderedNAT)) holds g . i in (Carrier (n --> OrderedNAT)) . i ) ) thus x = g ; ::_thesis: ( dom g = dom (Carrier (n --> OrderedNAT)) & ( for i being set st i in dom (Carrier (n --> OrderedNAT)) holds g . i in (Carrier (n --> OrderedNAT)) . i ) ) thus dom g = dom (Carrier (n --> OrderedNAT)) by A12, PARTFUN1:def_2; ::_thesis: for i being set st i in dom (Carrier (n --> OrderedNAT)) holds g . i in (Carrier (n --> OrderedNAT)) . i let i be set ; ::_thesis: ( i in dom (Carrier (n --> OrderedNAT)) implies g . i in (Carrier (n --> OrderedNAT)) . i ) assume i in dom (Carrier (n --> OrderedNAT)) ; ::_thesis: g . i in (Carrier (n --> OrderedNAT)) . i then A13: i in n ; then consider R being 1-sorted such that A14: R = (n --> OrderedNAT) . i and A15: (Carrier (n --> OrderedNAT)) . i = the carrier of R by PRALG_1:def_13; R = OrderedNAT by A13, A14, FUNCOP_1:7; hence g . i in (Carrier (n --> OrderedNAT)) . i by A15, DICKSON:def_15; ::_thesis: verum end; then x in product (Carrier (n --> OrderedNAT)) by CARD_3:def_5; hence x in the carrier of (product (n --> OrderedNAT)) by YELLOW_1:def_4; ::_thesis: verum end; hence the carrier of (product (n --> OrderedNAT)) = Bags n by TARSKI:1; ::_thesis: verum end; theorem Th33: :: BAGORDER:33 for n being Nat holds NaivelyOrderedBags n = product (n --> OrderedNAT) proof let n be Nat; ::_thesis: NaivelyOrderedBags n = product (n --> OrderedNAT) reconsider n0 = n as Element of NAT by ORDINAL1:def_12; set CNOB = the carrier of (NaivelyOrderedBags n); set IROB = the InternalRel of (NaivelyOrderedBags n); set CP = the carrier of (product (n --> OrderedNAT)); set IP = the InternalRel of (product (n --> OrderedNAT)); the carrier of (NaivelyOrderedBags n) = Bags n by Def11; then A1: the carrier of (NaivelyOrderedBags n) = the carrier of (product (n --> OrderedNAT)) by Th32; now__::_thesis:_for_a,_b_being_set_holds_ (_(_[a,b]_in_the_InternalRel_of_(NaivelyOrderedBags_n)_implies_[a,b]_in_the_InternalRel_of_(product_(n_-->_OrderedNAT))_)_&_(_[a,b]_in_the_InternalRel_of_(product_(n_-->_OrderedNAT))_implies_[a,b]_in_the_InternalRel_of_(NaivelyOrderedBags_n)_)_) let a, b be set ; ::_thesis: ( ( [a,b] in the InternalRel of (NaivelyOrderedBags n) implies [a,b] in the InternalRel of (product (n --> OrderedNAT)) ) & ( [a,b] in the InternalRel of (product (n --> OrderedNAT)) implies [a,b] in the InternalRel of (NaivelyOrderedBags n) ) ) hereby ::_thesis: ( [a,b] in the InternalRel of (product (n --> OrderedNAT)) implies [a,b] in the InternalRel of (NaivelyOrderedBags n) ) assume A2: [a,b] in the InternalRel of (NaivelyOrderedBags n) ; ::_thesis: [a,b] in the InternalRel of (product (n --> OrderedNAT)) then A3: a in dom the InternalRel of (NaivelyOrderedBags n) by XTUPLE_0:def_12; A4: b in rng the InternalRel of (NaivelyOrderedBags n) by A2, XTUPLE_0:def_13; A5: a in the carrier of (NaivelyOrderedBags n) by A3; A6: b in the carrier of (NaivelyOrderedBags n) by A4; A7: a in Bags n by A5, Def11; A8: b in Bags n by A6, Def11; then reconsider a9 = a, b9 = b as Element of the carrier of (product (n --> OrderedNAT)) by A7, Th32; a9 in the carrier of (product (n0 --> OrderedNAT)) ; then A9: a9 in product (Carrier (n --> OrderedNAT)) by YELLOW_1:def_4; now__::_thesis:_ex_f,_g_being_Function_st_ (_f_=_a9_&_g_=_b9_&_(_for_i_being_set_st_i_in_n_holds_ ex_R_being_RelStr_ex_xi,_yi_being_Element_of_R_st_ (_R_=_(n_-->_OrderedNAT)_._i_&_xi_=_f_._i_&_yi_=_g_._i_&_xi_<=_yi_)_)_) set f = a9; set g = b9; A10: a9 is bag of n by A7; A11: b is bag of n by A8; reconsider f = a9, g = b9 as Function by A7, A8; take f = f; ::_thesis: ex g being Function st ( f = a9 & g = b9 & ( for i being set st i in n holds ex R being RelStr ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) ) ) take g = g; ::_thesis: ( f = a9 & g = b9 & ( for i being set st i in n holds ex R being RelStr ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) ) ) thus ( f = a9 & g = b9 ) ; ::_thesis: for i being set st i in n holds ex R being RelStr ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) let i be set ; ::_thesis: ( i in n implies ex R being RelStr ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) ) assume A12: i in n ; ::_thesis: ex R being RelStr ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) set R = (n --> OrderedNAT) . i; A13: (n --> OrderedNAT) . i = OrderedNAT by A12, FUNCOP_1:7; reconsider R = (n --> OrderedNAT) . i as RelStr by A12, FUNCOP_1:7; take R = R; ::_thesis: ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) set xi = f . i; set yi = g . i; dom f = n by A10, PARTFUN1:def_2; then A14: f . i in rng f by A12, FUNCT_1:3; A15: rng f c= NAT by A10, VALUED_0:def_6; dom g = n by A11, PARTFUN1:def_2; then A16: g . i in rng g by A12, FUNCT_1:3; rng g c= NAT by A11, VALUED_0:def_6; then reconsider xi = f . i, yi = g . i as Element of R by A12, A14, A15, A16, DICKSON:def_15, FUNCOP_1:7; take xi = xi; ::_thesis: ex yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) take yi = yi; ::_thesis: ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) thus ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i ) ; ::_thesis: xi <= yi reconsider a99 = a9, b99 = b9 as bag of n by A7, A8; a99 divides b99 by A2, Def11; then a99 . i <= b99 . i by PRE_POLY:def_11; then [xi,yi] in NATOrd by DICKSON:def_14; hence xi <= yi by A13, DICKSON:def_15, ORDERS_2:def_5; ::_thesis: verum end; then a9 <= b9 by A9, YELLOW_1:def_4; hence [a,b] in the InternalRel of (product (n --> OrderedNAT)) by ORDERS_2:def_5; ::_thesis: verum end; assume A17: [a,b] in the InternalRel of (product (n --> OrderedNAT)) ; ::_thesis: [a,b] in the InternalRel of (NaivelyOrderedBags n) then A18: a in dom the InternalRel of (product (n --> OrderedNAT)) by XTUPLE_0:def_12; A19: b in rng the InternalRel of (product (n --> OrderedNAT)) by A17, XTUPLE_0:def_13; A20: a in the carrier of (product (n --> OrderedNAT)) by A18; A21: b in the carrier of (product (n --> OrderedNAT)) by A19; A22: a in Bags n by A20, Th32; b in Bags n by A21, Th32; then reconsider a9 = a, b9 = b as bag of n by A22; reconsider a2 = a9, b2 = b9 as Element of the carrier of (product (n --> OrderedNAT)) by A18, A19; a2 in the carrier of (product (n0 --> OrderedNAT)) ; then A23: a2 in product (Carrier (n --> OrderedNAT)) by YELLOW_1:def_4; a2 <= b2 by A17, ORDERS_2:def_5; then consider f, g being Function such that A24: f = a2 and A25: g = b2 and A26: for i being set st i in n holds ex R being RelStr ex xi, yi being Element of R st ( R = (n --> OrderedNAT) . i & xi = f . i & yi = g . i & xi <= yi ) by A23, YELLOW_1:def_4; now__::_thesis:_for_k_being_set_st_k_in_n_holds_ a9_._k_<=_b9_._k let k be set ; ::_thesis: ( k in n implies a9 . k <= b9 . k ) assume A27: k in n ; ::_thesis: a9 . k <= b9 . k consider R being RelStr , xi, yi being Element of R such that A28: R = (n --> OrderedNAT) . k and A29: xi = f . k and A30: yi = g . k and A31: xi <= yi by A26, A27; R = OrderedNAT by A27, A28, FUNCOP_1:7; then [xi,yi] in NATOrd by A31, DICKSON:def_15, ORDERS_2:def_5; then consider xii, yii being Element of NAT such that A32: [xii,yii] = [xi,yi] and A33: xii <= yii by DICKSON:def_14; xii = xi by A32, XTUPLE_0:1; hence a9 . k <= b9 . k by A24, A25, A29, A30, A32, A33, XTUPLE_0:1; ::_thesis: verum end; then a9 divides b9 by PRE_POLY:46; hence [a,b] in the InternalRel of (NaivelyOrderedBags n) by Def11; ::_thesis: verum end; hence NaivelyOrderedBags n = product (n --> OrderedNAT) by A1, RELAT_1:def_2; ::_thesis: verum end; theorem :: BAGORDER:34 for n being Nat for o being TermOrder of n st o is admissible holds ( the InternalRel of (NaivelyOrderedBags n) c= o & o is well-ordering ) proof let n be Nat; ::_thesis: for o being TermOrder of n st o is admissible holds ( the InternalRel of (NaivelyOrderedBags n) c= o & o is well-ordering ) let o be TermOrder of n; ::_thesis: ( o is admissible implies ( the InternalRel of (NaivelyOrderedBags n) c= o & o is well-ordering ) ) assume A1: o is admissible ; ::_thesis: ( the InternalRel of (NaivelyOrderedBags n) c= o & o is well-ordering ) reconsider n0 = n as Element of NAT by ORDINAL1:def_12; set IRN = the InternalRel of (NaivelyOrderedBags n); now__::_thesis:_for_a,_b_being_set_st_[a,b]_in_the_InternalRel_of_(NaivelyOrderedBags_n)_holds_ [a,b]_in_o let a, b be set ; ::_thesis: ( [a,b] in the InternalRel of (NaivelyOrderedBags n) implies [a,b] in o ) assume A2: [a,b] in the InternalRel of (NaivelyOrderedBags n) ; ::_thesis: [a,b] in o A3: a in dom the InternalRel of (NaivelyOrderedBags n) by A2, XTUPLE_0:def_12; b in rng the InternalRel of (NaivelyOrderedBags n) by A2, XTUPLE_0:def_13; then reconsider a1 = a, b1 = b as Element of Bags n by A3, Def11; A4: a1 divides b1 by A2, Def11; set l = b1 -' a1; now__::_thesis:_for_i_being_set_st_i_in_n_holds_ ((b1_-'_a1)_+_a1)_._i_=_b1_._i let i be set ; ::_thesis: ( i in n implies ((b1 -' a1) + a1) . i = b1 . i ) assume i in n ; ::_thesis: ((b1 -' a1) + a1) . i = b1 . i A5: ((b1 -' a1) + a1) . i = ((b1 -' a1) . i) + (a1 . i) by PRE_POLY:def_5 .= ((b1 . i) -' (a1 . i)) + (a1 . i) by PRE_POLY:def_6 ; a1 . i <= b1 . i by A4, PRE_POLY:def_11; then (a1 . i) - (a1 . i) <= (b1 . i) - (a1 . i) by XREAL_1:9; hence ((b1 -' a1) + a1) . i = ((b1 . i) + (- (a1 . i))) + (a1 . i) by A5, XREAL_0:def_2 .= b1 . i ; ::_thesis: verum end; then A6: (b1 -' a1) + a1 = b1 by PBOOLE:3; [(EmptyBag n),(b1 -' a1)] in o by A1, Def5; then [((EmptyBag n) + a1),((b1 -' a1) + a1)] in o by A1, Def5; hence [a,b] in o by A6, PRE_POLY:53; ::_thesis: verum end; hence A7: the InternalRel of (NaivelyOrderedBags n) c= o by RELAT_1:def_3; ::_thesis: o is well-ordering set R = product (n0 --> OrderedNAT); set S = RelStr(# (Bags n),o #); A8: RelStr(# (Bags n),o #) is quasi_ordered by DICKSON:def_3; A9: the InternalRel of (product (n0 --> OrderedNAT)) c= the InternalRel of RelStr(# (Bags n),o #) by A7, Th33; the carrier of (product (n0 --> OrderedNAT)) = the carrier of RelStr(# (Bags n),o #) by Th32; then A10: RelStr(# (Bags n),o #) \~ is well_founded by A8, A9, DICKSON:49; o is_strongly_connected_in Bags n by A1, Def5; then A11: o is_connected_in Bags n by ORDERS_1:7; A12: field o = Bags n by ORDERS_1:12; RelStr(# (Bags n),o #) is well_founded by A10, DICKSON:16; then A13: o is_well_founded_in field o by A12, WELLFND1:def_2; A14: o is connected by A11, A12, RELAT_2:def_14; o is well_founded by A13, WELLORD1:3; hence o is well-ordering by A14; ::_thesis: verum end; begin definition let R be non empty connected Poset; let X be Element of Fin the carrier of R; assume B1: not X is empty ; func PosetMin X -> Element of R means :: BAGORDER:def 12 ( it in X & it is_minimal_wrt X, the InternalRel of R ); existence ex b1 being Element of R st ( b1 in X & b1 is_minimal_wrt X, the InternalRel of R ) proof set IR = the InternalRel of R; X c= the carrier of R by FINSUB_1:def_5; then consider x being Element of R such that A1: x in X and A2: x is_minimal_wrt X, the InternalRel of R by B1, Th7; take x ; ::_thesis: ( x in X & x is_minimal_wrt X, the InternalRel of R ) thus ( x in X & x is_minimal_wrt X, the InternalRel of R ) by A1, A2; ::_thesis: verum end; uniqueness for b1, b2 being Element of R st b1 in X & b1 is_minimal_wrt X, the InternalRel of R & b2 in X & b2 is_minimal_wrt X, the InternalRel of R holds b1 = b2 proof let IT1, IT2 be Element of R; ::_thesis: ( IT1 in X & IT1 is_minimal_wrt X, the InternalRel of R & IT2 in X & IT2 is_minimal_wrt X, the InternalRel of R implies IT1 = IT2 ) assume that A3: IT1 in X and A4: IT1 is_minimal_wrt X, the InternalRel of R and A5: IT2 in X and A6: IT2 is_minimal_wrt X, the InternalRel of R ; ::_thesis: IT1 = IT2 set IR = the InternalRel of R; A7: ( IT1 <= IT2 or IT2 <= IT1 ) by WAYBEL_0:def_29; percases ( [IT1,IT2] in the InternalRel of R or [IT2,IT1] in the InternalRel of R ) by A7, ORDERS_2:def_5; suppose [IT1,IT2] in the InternalRel of R ; ::_thesis: IT1 = IT2 hence IT1 = IT2 by A3, A6, WAYBEL_4:def_25; ::_thesis: verum end; suppose [IT2,IT1] in the InternalRel of R ; ::_thesis: IT1 = IT2 hence IT1 = IT2 by A4, A5, WAYBEL_4:def_25; ::_thesis: verum end; end; end; func PosetMax X -> Element of R means :Def13: :: BAGORDER:def 13 ( it in X & it is_maximal_wrt X, the InternalRel of R ); existence ex b1 being Element of R st ( b1 in X & b1 is_maximal_wrt X, the InternalRel of R ) proof set IR = the InternalRel of R; X c= the carrier of R by FINSUB_1:def_5; then consider x being Element of R such that A8: x in X and A9: x is_maximal_wrt X, the InternalRel of R by B1, Th6; take x ; ::_thesis: ( x in X & x is_maximal_wrt X, the InternalRel of R ) thus ( x in X & x is_maximal_wrt X, the InternalRel of R ) by A8, A9; ::_thesis: verum end; uniqueness for b1, b2 being Element of R st b1 in X & b1 is_maximal_wrt X, the InternalRel of R & b2 in X & b2 is_maximal_wrt X, the InternalRel of R holds b1 = b2 proof let IT1, IT2 be Element of R; ::_thesis: ( IT1 in X & IT1 is_maximal_wrt X, the InternalRel of R & IT2 in X & IT2 is_maximal_wrt X, the InternalRel of R implies IT1 = IT2 ) assume that A10: IT1 in X and A11: IT1 is_maximal_wrt X, the InternalRel of R and A12: IT2 in X and A13: IT2 is_maximal_wrt X, the InternalRel of R ; ::_thesis: IT1 = IT2 set IR = the InternalRel of R; A14: ( IT1 <= IT2 or IT2 <= IT1 ) by WAYBEL_0:def_29; percases ( [IT1,IT2] in the InternalRel of R or [IT2,IT1] in the InternalRel of R ) by A14, ORDERS_2:def_5; suppose [IT1,IT2] in the InternalRel of R ; ::_thesis: IT1 = IT2 hence IT1 = IT2 by A11, A12, WAYBEL_4:def_23; ::_thesis: verum end; suppose [IT2,IT1] in the InternalRel of R ; ::_thesis: IT1 = IT2 hence IT1 = IT2 by A10, A13, WAYBEL_4:def_23; ::_thesis: verum end; end; end; end; :: deftheorem defines PosetMin BAGORDER:def_12_:_ for R being non empty connected Poset for X being Element of Fin the carrier of R st not X is empty holds for b3 being Element of R holds ( b3 = PosetMin X iff ( b3 in X & b3 is_minimal_wrt X, the InternalRel of R ) ); :: deftheorem Def13 defines PosetMax BAGORDER:def_13_:_ for R being non empty connected Poset for X being Element of Fin the carrier of R st not X is empty holds for b3 being Element of R holds ( b3 = PosetMax X iff ( b3 in X & b3 is_maximal_wrt X, the InternalRel of R ) ); definition let R be non empty connected Poset; func FinOrd-Approx R -> Function of NAT,(bool [:(Fin the carrier of R),(Fin the carrier of R):]) means :Def14: :: BAGORDER:def 14 ( dom it = NAT & it . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds it . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in it . n ) } ) ); existence ex b1 being Function of NAT,(bool [:(Fin the carrier of R),(Fin the carrier of R):]) st ( dom b1 = NAT & b1 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds b1 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in b1 . n ) } ) ) proof set IR = the InternalRel of R; set CR = the carrier of R; set FBCP = [:(Fin the carrier of R),(Fin the carrier of R):]; defpred S1[ Nat, set , set ] means $3 = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in $2 ) } ; A1: for n being Element of NAT for x being set ex y being set st S1[n,x,y] ; consider f being Function such that A2: dom f = NAT and A3: f . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } and A4: for n being Element of NAT holds S1[n,f . n,f . (n + 1)] from RECDEF_1:sch_1(A1); A5: for n being Nat holds S1[n,f . n,f . (n + 1)] proof let n be Nat; ::_thesis: S1[n,f . n,f . (n + 1)] n in NAT by ORDINAL1:def_12; hence S1[n,f . n,f . (n + 1)] by A4; ::_thesis: verum end; now__::_thesis:_(_dom_f_=_NAT_&_(_for_x_being_set_st_x_in_NAT_holds_ f_._x_in_bool_[:(Fin_the_carrier_of_R),(Fin_the_carrier_of_R):]_)_) thus dom f = NAT by A2; ::_thesis: for x being set st x in NAT holds f . b2 in bool [:(Fin the carrier of R),(Fin the carrier of R):] let x be set ; ::_thesis: ( x in NAT implies f . b1 in bool [:(Fin the carrier of R),(Fin the carrier of R):] ) assume x in NAT ; ::_thesis: f . b1 in bool [:(Fin the carrier of R),(Fin the carrier of R):] then reconsider x9 = x as Element of NAT ; percases ( x9 = 0 or x9 > 0 ) ; supposeA6: x9 = 0 ; ::_thesis: f . b1 in bool [:(Fin the carrier of R),(Fin the carrier of R):] set F0 = { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } ; now__::_thesis:_for_z_being_set_st_z_in__{__[a,b]_where_a,_b_is_Element_of_Fin_the_carrier_of_R_:_(_a_=_{}_or_(_a_<>_{}_&_b_<>_{}_&_PosetMax_a_<>_PosetMax_b_&_[(PosetMax_a),(PosetMax_b)]_in_the_InternalRel_of_R_)_)__}__holds_ z_in_[:(Fin_the_carrier_of_R),(Fin_the_carrier_of_R):] let z be set ; ::_thesis: ( z in { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } implies z in [:(Fin the carrier of R),(Fin the carrier of R):] ) assume z in { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } ; ::_thesis: z in [:(Fin the carrier of R),(Fin the carrier of R):] then ex a, b being Element of Fin the carrier of R st ( z = [a,b] & ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) ) ; hence z in [:(Fin the carrier of R),(Fin the carrier of R):] ; ::_thesis: verum end; then { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } c= [:(Fin the carrier of R),(Fin the carrier of R):] by TARSKI:def_3; hence f . x in bool [:(Fin the carrier of R),(Fin the carrier of R):] by A3, A6; ::_thesis: verum end; supposeA7: x9 > 0 ; ::_thesis: f . b1 in bool [:(Fin the carrier of R),(Fin the carrier of R):] A8: x9 = (x9 - 1) + 1 ; reconsider x1 = x9 - 1 as Element of NAT by A7, NAT_1:20; set FX = { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in f . x1 ) } ; A9: { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in f . x1 ) } = f . x9 by A4, A8; now__::_thesis:_for_z_being_set_st_z_in__{__[a,b]_where_a,_b_is_Element_of_Fin_the_carrier_of_R_:_(_a_<>_{}_&_b_<>_{}_&_PosetMax_a_=_PosetMax_b_&_[(a_\_{(PosetMax_a)}),(b_\_{(PosetMax_b)})]_in_f_._x1_)__}__holds_ z_in_[:(Fin_the_carrier_of_R),(Fin_the_carrier_of_R):] let z be set ; ::_thesis: ( z in { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in f . x1 ) } implies z in [:(Fin the carrier of R),(Fin the carrier of R):] ) assume z in { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in f . x1 ) } ; ::_thesis: z in [:(Fin the carrier of R),(Fin the carrier of R):] then ex a, b being Element of Fin the carrier of R st ( z = [a,b] & a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in f . x1 ) ; hence z in [:(Fin the carrier of R),(Fin the carrier of R):] ; ::_thesis: verum end; then { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in f . x1 ) } c= [:(Fin the carrier of R),(Fin the carrier of R):] by TARSKI:def_3; hence f . x in bool [:(Fin the carrier of R),(Fin the carrier of R):] by A9; ::_thesis: verum end; end; end; then reconsider f = f as Function of NAT,(bool [:(Fin the carrier of R),(Fin the carrier of R):]) by FUNCT_2:3; take f ; ::_thesis: ( dom f = NAT & f . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds f . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in f . n ) } ) ) thus ( dom f = NAT & f . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds f . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in f . n ) } ) ) by A2, A3, A5; ::_thesis: verum end; uniqueness for b1, b2 being Function of NAT,(bool [:(Fin the carrier of R),(Fin the carrier of R):]) st dom b1 = NAT & b1 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds b1 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in b1 . n ) } ) & dom b2 = NAT & b2 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds b2 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in b2 . n ) } ) holds b1 = b2 proof set IR = the InternalRel of R; set CR = the carrier of R; set FBCP = [:(Fin the carrier of R),(Fin the carrier of R):]; let IT1, IT2 be Function of NAT,(bool [:(Fin the carrier of R),(Fin the carrier of R):]); ::_thesis: ( dom IT1 = NAT & IT1 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds IT1 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in IT1 . n ) } ) & dom IT2 = NAT & IT2 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds IT2 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in IT2 . n ) } ) implies IT1 = IT2 ) assume that A10: dom IT1 = NAT and A11: IT1 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } and A12: for n being Nat holds IT1 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in IT1 . n ) } and A13: dom IT2 = NAT and A14: IT2 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } and A15: for n being Nat holds IT2 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in IT2 . n ) } ; ::_thesis: IT1 = IT2 defpred S1[ Nat, set , set ] means $3 = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in $2 ) } ; A16: dom IT1 = NAT by A10; A17: IT1 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } by A11; A18: for n being Nat holds S1[n,IT1 . n,IT1 . (n + 1)] by A12; A19: dom IT2 = NAT by A13; A20: IT2 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } by A14; A21: for n being Nat holds S1[n,IT2 . n,IT2 . (n + 1)] by A15; A22: for n being Nat for x, y1, y2 being set st S1[n,x,y1] & S1[n,x,y2] holds y1 = y2 ; thus IT1 = IT2 from NAT_1:sch_13(A16, A17, A18, A19, A20, A21, A22); ::_thesis: verum end; end; :: deftheorem Def14 defines FinOrd-Approx BAGORDER:def_14_:_ for R being non empty connected Poset for b2 being Function of NAT,(bool [:(Fin the carrier of R),(Fin the carrier of R):]) holds ( b2 = FinOrd-Approx R iff ( dom b2 = NAT & b2 . 0 = { [x,y] where x, y is Element of Fin the carrier of R : ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ) } & ( for n being Nat holds b2 . (n + 1) = { [x,y] where x, y is Element of Fin the carrier of R : ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in b2 . n ) } ) ) ); theorem Th35: :: BAGORDER:35 for R being non empty connected Poset for x, y being Element of Fin the carrier of R holds ( [x,y] in union (rng (FinOrd-Approx R)) iff ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) ) proof let R be non empty connected Poset; ::_thesis: for x, y being Element of Fin the carrier of R holds ( [x,y] in union (rng (FinOrd-Approx R)) iff ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) ) let x, y be Element of Fin the carrier of R; ::_thesis: ( [x,y] in union (rng (FinOrd-Approx R)) iff ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) ) set IR = the InternalRel of R; set CR = the carrier of R; set FOAR = FinOrd-Approx R; A1: (FinOrd-Approx R) . 0 = { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } by Def14; A2: dom (FinOrd-Approx R) = NAT by Def14; hereby ::_thesis: ( ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) implies [x,y] in union (rng (FinOrd-Approx R)) ) assume [x,y] in union (rng (FinOrd-Approx R)) ; ::_thesis: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) then consider Y being set such that A3: [x,y] in Y and A4: Y in rng (FinOrd-Approx R) by TARSKI:def_4; consider n being set such that A5: n in dom (FinOrd-Approx R) and A6: Y = (FinOrd-Approx R) . n by A4, FUNCT_1:def_3; reconsider n = n as Element of NAT by A5; percases ( n = 0 or n > 0 ) ; suppose n = 0 ; ::_thesis: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) then consider a, b being Element of Fin the carrier of R such that A7: [x,y] = [a,b] and A8: ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) by A1, A3, A6; x = a by A7, XTUPLE_0:1; hence ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) by A7, A8, XTUPLE_0:1; ::_thesis: verum end; suppose n > 0 ; ::_thesis: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) then A9: n - 1 is Element of NAT by NAT_1:20; then (FinOrd-Approx R) . ((n - 1) + 1) = { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in (FinOrd-Approx R) . (n - 1) ) } by Def14; then consider a, b being Element of Fin the carrier of R such that A10: [x,y] = [a,b] and a <> {} and A11: b <> {} and A12: PosetMax a = PosetMax b and A13: [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in (FinOrd-Approx R) . (n - 1) by A3, A6; A14: x = a by A10, XTUPLE_0:1; A15: y = b by A10, XTUPLE_0:1; (FinOrd-Approx R) . (n - 1) in rng (FinOrd-Approx R) by A2, A9, FUNCT_1:def_3; hence ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) by A11, A12, A13, A14, A15, TARSKI:def_4; ::_thesis: verum end; end; end; assume A16: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) ; ::_thesis: [x,y] in union (rng (FinOrd-Approx R)) percases ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ) by A16; suppose x = {} ; ::_thesis: [x,y] in union (rng (FinOrd-Approx R)) then A17: [x,y] in (FinOrd-Approx R) . 0 by A1; (FinOrd-Approx R) . 0 in rng (FinOrd-Approx R) by A2, FUNCT_1:def_3; hence [x,y] in union (rng (FinOrd-Approx R)) by A17, TARSKI:def_4; ::_thesis: verum end; suppose ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ; ::_thesis: [x,y] in union (rng (FinOrd-Approx R)) then A18: [x,y] in (FinOrd-Approx R) . 0 by A1; (FinOrd-Approx R) . 0 in rng (FinOrd-Approx R) by A2, FUNCT_1:def_3; hence [x,y] in union (rng (FinOrd-Approx R)) by A18, TARSKI:def_4; ::_thesis: verum end; supposeA19: ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in union (rng (FinOrd-Approx R)) ) ; ::_thesis: [x,y] in union (rng (FinOrd-Approx R)) set NEXTXY = [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})]; consider Y being set such that A20: [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in Y and A21: Y in rng (FinOrd-Approx R) by A19, TARSKI:def_4; consider n being set such that A22: n in dom (FinOrd-Approx R) and A23: Y = (FinOrd-Approx R) . n by A21, FUNCT_1:def_3; reconsider n = n as Nat by A22; (FinOrd-Approx R) . (n + 1) = { [a,b] where a, b is Element of Fin the carrier of R : ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in (FinOrd-Approx R) . n ) } by Def14; then A24: [x,y] in (FinOrd-Approx R) . (n + 1) by A19, A20, A23; (FinOrd-Approx R) . (n + 1) in rng (FinOrd-Approx R) by A2, FUNCT_1:def_3; hence [x,y] in union (rng (FinOrd-Approx R)) by A24, TARSKI:def_4; ::_thesis: verum end; end; end; theorem Th36: :: BAGORDER:36 for R being non empty connected Poset for x being Element of Fin the carrier of R st x <> {} holds not [x,{}] in union (rng (FinOrd-Approx R)) proof let R be non empty connected Poset; ::_thesis: for x being Element of Fin the carrier of R st x <> {} holds not [x,{}] in union (rng (FinOrd-Approx R)) let x be Element of Fin the carrier of R; ::_thesis: ( x <> {} implies not [x,{}] in union (rng (FinOrd-Approx R)) ) assume A1: x <> {} ; ::_thesis: not [x,{}] in union (rng (FinOrd-Approx R)) set CR = the carrier of R; set FOAR = FinOrd-Approx R; reconsider y = {} as Element of Fin the carrier of R by FINSUB_1:7; now__::_thesis:_not_[x,y]_in_union_(rng_(FinOrd-Approx_R)) assume A2: [x,y] in union (rng (FinOrd-Approx R)) ; ::_thesis: contradiction percases ( x = {} or ( x <> {} & y <> {} & [(PosetMax x),(PosetMax y)] in the carrier of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ (PosetMax x)),({} \ (PosetMax y))] in union (rng (FinOrd-Approx R)) ) ) by A2, Th35; suppose x = {} ; ::_thesis: contradiction hence contradiction by A1; ::_thesis: verum end; suppose ( x <> {} & y <> {} & [(PosetMax x),(PosetMax y)] in the carrier of R ) ; ::_thesis: contradiction hence contradiction ; ::_thesis: verum end; suppose ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ (PosetMax x)),({} \ (PosetMax y))] in union (rng (FinOrd-Approx R)) ) ; ::_thesis: contradiction hence contradiction ; ::_thesis: verum end; end; end; hence not [x,{}] in union (rng (FinOrd-Approx R)) ; ::_thesis: verum end; theorem Th37: :: BAGORDER:37 for R being non empty connected Poset for a being Element of Fin the carrier of R holds a \ {(PosetMax a)} is Element of Fin the carrier of R proof let R be non empty connected Poset; ::_thesis: for a being Element of Fin the carrier of R holds a \ {(PosetMax a)} is Element of Fin the carrier of R let a be Element of Fin the carrier of R; ::_thesis: a \ {(PosetMax a)} is Element of Fin the carrier of R set CR = the carrier of R; A1: a c= the carrier of R by FINSUB_1:def_5; reconsider a9 = a as finite set ; set z = a9 \ {(PosetMax a)}; a9 \ {(PosetMax a)} c= the carrier of R by A1, XBOOLE_1:1; hence a \ {(PosetMax a)} is Element of Fin the carrier of R by FINSUB_1:def_5; ::_thesis: verum end; Lm1: for R being non empty connected Poset for n being Nat for a being Element of Fin the carrier of R st card a = n + 1 holds card (a \ {(PosetMax a)}) = n proof let R be non empty connected Poset; ::_thesis: for n being Nat for a being Element of Fin the carrier of R st card a = n + 1 holds card (a \ {(PosetMax a)}) = n let n be Nat; ::_thesis: for a being Element of Fin the carrier of R st card a = n + 1 holds card (a \ {(PosetMax a)}) = n let a be Element of Fin the carrier of R; ::_thesis: ( card a = n + 1 implies card (a \ {(PosetMax a)}) = n ) assume A1: card a = n + 1 ; ::_thesis: card (a \ {(PosetMax a)}) = n reconsider a9 = a as finite set ; now__::_thesis:_for_w_being_set_st_w_in_{(PosetMax_a)}_holds_ w_in_a let w be set ; ::_thesis: ( w in {(PosetMax a)} implies w in a ) assume w in {(PosetMax a)} ; ::_thesis: w in a then w = PosetMax a by TARSKI:def_1; hence w in a by A1, Def13, CARD_1:27; ::_thesis: verum end; then {(PosetMax a)} c= a by TARSKI:def_3; then A2: card (a9 \ {(PosetMax a)}) = (card a9) - (card {(PosetMax a)}) by CARD_2:44; card {(PosetMax a)} = 1 by CARD_1:30; hence card (a \ {(PosetMax a)}) = n by A1, A2; ::_thesis: verum end; theorem Th38: :: BAGORDER:38 for R being non empty connected Poset holds union (rng (FinOrd-Approx R)) is Order of (Fin the carrier of R) proof let R be non empty connected Poset; ::_thesis: union (rng (FinOrd-Approx R)) is Order of (Fin the carrier of R) set IR = the InternalRel of R; set CR = the carrier of R; set X = union (rng (FinOrd-Approx R)); set FOAR = FinOrd-Approx R; set FOAR0 = { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } ; A1: (FinOrd-Approx R) . 0 = { [a,b] where a, b is Element of Fin the carrier of R : ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ) } by Def14; now__::_thesis:_for_x_being_set_st_x_in_union_(rng_(FinOrd-Approx_R))_holds_ x_in_[:(Fin_the_carrier_of_R),(Fin_the_carrier_of_R):] let x be set ; ::_thesis: ( x in union (rng (FinOrd-Approx R)) implies x in [:(Fin the carrier of R),(Fin the carrier of R):] ) assume x in union (rng (FinOrd-Approx R)) ; ::_thesis: x in [:(Fin the carrier of R),(Fin the carrier of R):] then A2: ex Y being set st ( x in Y & Y in rng (FinOrd-Approx R) ) by TARSKI:def_4; rng (FinOrd-Approx R) c= bool [:(Fin the carrier of R),(Fin the carrier of R):] by RELAT_1:def_19; hence x in [:(Fin the carrier of R),(Fin the carrier of R):] by A2; ::_thesis: verum end; then reconsider X = union (rng (FinOrd-Approx R)) as Relation of (Fin the carrier of R) by TARSKI:def_3; A3: now__::_thesis:_for_x_being_set_st_x_in_Fin_the_carrier_of_R_holds_ [x,x]_in_X let x be set ; ::_thesis: ( x in Fin the carrier of R implies [x,x] in X ) assume A4: x in Fin the carrier of R ; ::_thesis: [x,x] in X 0 in NAT ; then 0 in dom (FinOrd-Approx R) by Def14; then A5: (FinOrd-Approx R) . 0 in rng (FinOrd-Approx R) by FUNCT_1:def_3; reconsider x9 = x as Element of Fin the carrier of R by A4; defpred S1[ Nat] means for x being Element of Fin the carrier of R st card x = $1 holds [x,x] in union (rng (FinOrd-Approx R)); A6: S1[ 0 ] proof let x be Element of Fin the carrier of R; ::_thesis: ( card x = 0 implies [x,x] in union (rng (FinOrd-Approx R)) ) assume card x = 0 ; ::_thesis: [x,x] in union (rng (FinOrd-Approx R)) then x = {} ; then [x,x] in (FinOrd-Approx R) . 0 by A1; hence [x,x] in union (rng (FinOrd-Approx R)) by A5, TARSKI:def_4; ::_thesis: verum end; A7: for n being Nat st S1[n] holds S1[n + 1] proof let n be Nat; ::_thesis: ( S1[n] implies S1[n + 1] ) assume A8: for x being Element of Fin the carrier of R st card x = n holds [x,x] in union (rng (FinOrd-Approx R)) ; ::_thesis: S1[n + 1] let y be Element of Fin the carrier of R; ::_thesis: ( card y = n + 1 implies [y,y] in union (rng (FinOrd-Approx R)) ) assume A9: card y = n + 1 ; ::_thesis: [y,y] in union (rng (FinOrd-Approx R)) percases ( y = {} or y <> {} ) ; suppose y = {} ; ::_thesis: [y,y] in union (rng (FinOrd-Approx R)) then [y,y] in (FinOrd-Approx R) . 0 by A1; hence [y,y] in union (rng (FinOrd-Approx R)) by A5, TARSKI:def_4; ::_thesis: verum end; supposeA10: y <> {} ; ::_thesis: [y,y] in union (rng (FinOrd-Approx R)) set z = y \ {(PosetMax y)}; reconsider z = y \ {(PosetMax y)} as Element of Fin the carrier of R by Th37; card z = n by A9, Lm1; then [z,z] in union (rng (FinOrd-Approx R)) by A8; hence [y,y] in union (rng (FinOrd-Approx R)) by A10, Th35; ::_thesis: verum end; end; end; A11: for n being Nat holds S1[n] from NAT_1:sch_2(A6, A7); consider n being Nat such that A12: x,n are_equipotent by A4, CARD_1:43; card x9 = n by A12, CARD_1:def_2; hence [x,x] in X by A11; ::_thesis: verum end; A13: now__::_thesis:_for_x,_y_being_set_st_x_in_Fin_the_carrier_of_R_&_y_in_Fin_the_carrier_of_R_&_[x,y]_in_X_&_[y,x]_in_X_holds_ x_=_y let x, y be set ; ::_thesis: ( x in Fin the carrier of R & y in Fin the carrier of R & [x,y] in X & [y,x] in X implies x = y ) assume that A14: x in Fin the carrier of R and A15: y in Fin the carrier of R and A16: [x,y] in X and A17: [y,x] in X ; ::_thesis: x = y reconsider x9 = x as Element of Fin the carrier of R by A14; defpred S1[ Nat] means for x, y being Element of Fin the carrier of R st card x = $1 & [x,y] in X & [y,x] in X holds x = y; now__::_thesis:_for_a,_b_being_Element_of_Fin_the_carrier_of_R_st_card_a_=_0_&_[a,b]_in_X_&_[b,a]_in_X_holds_ a_=_b let a, b be Element of Fin the carrier of R; ::_thesis: ( card a = 0 & [a,b] in X & [b,a] in X implies a = b ) assume that A18: card a = 0 and [a,b] in X and A19: [b,a] in X ; ::_thesis: a = b reconsider a9 = a as finite set ; a9 = {} by A18; hence a = b by A19, Th36; ::_thesis: verum end; then A20: S1[ 0 ] ; now__::_thesis:_for_n_being_Nat_st_(_for_a,_b_being_Element_of_Fin_the_carrier_of_R_st_card_a_=_n_&_[a,b]_in_X_&_[b,a]_in_X_holds_ a_=_b_)_holds_ for_a,_b_being_Element_of_Fin_the_carrier_of_R_st_card_a_=_n_+_1_&_[a,b]_in_X_&_[b,a]_in_X_holds_ a_=_b let n be Nat; ::_thesis: ( ( for a, b being Element of Fin the carrier of R st card a = n & [a,b] in X & [b,a] in X holds a = b ) implies for a, b being Element of Fin the carrier of R st card a = n + 1 & [a,b] in X & [b,a] in X holds b4 = b5 ) assume A21: for a, b being Element of Fin the carrier of R st card a = n & [a,b] in X & [b,a] in X holds a = b ; ::_thesis: for a, b being Element of Fin the carrier of R st card a = n + 1 & [a,b] in X & [b,a] in X holds b4 = b5 let a, b be Element of Fin the carrier of R; ::_thesis: ( card a = n + 1 & [a,b] in X & [b,a] in X implies b2 = b3 ) assume that A22: card a = n + 1 and A23: [a,b] in X and A24: [b,a] in X ; ::_thesis: b2 = b3 percases ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) or ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in X ) ) by A23, Th35; suppose a = {} ; ::_thesis: b2 = b3 hence a = b by A22; ::_thesis: verum end; supposeA25: ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ; ::_thesis: b2 = b3 now__::_thesis:_a_=_b percases ( b = {} or ( b <> {} & a <> {} & PosetMax b <> PosetMax a & [(PosetMax b),(PosetMax a)] in the InternalRel of R ) or ( b <> {} & a <> {} & PosetMax b = PosetMax a & [(b \ {(PosetMax b)}),(a \ {(PosetMax a)})] in X ) ) by A24, Th35; suppose b = {} ; ::_thesis: a = b hence a = b by A25; ::_thesis: verum end; suppose ( b <> {} & a <> {} & PosetMax b <> PosetMax a & [(PosetMax b),(PosetMax a)] in the InternalRel of R ) ; ::_thesis: a = b hence a = b by A25, ORDERS_1:4; ::_thesis: verum end; suppose ( b <> {} & a <> {} & PosetMax b = PosetMax a & [(b \ {(PosetMax b)}),(a \ {(PosetMax a)})] in X ) ; ::_thesis: a = b hence a = b by A25; ::_thesis: verum end; end; end; hence a = b ; ::_thesis: verum end; supposeA26: ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in X ) ; ::_thesis: b2 = b3 now__::_thesis:_a_=_b percases ( b = {} or ( b <> {} & a <> {} & PosetMax b <> PosetMax a & [(PosetMax b),(PosetMax a)] in the InternalRel of R ) or ( b <> {} & a <> {} & PosetMax b = PosetMax a & [(b \ {(PosetMax b)}),(a \ {(PosetMax a)})] in X ) ) by A24, Th35; suppose b = {} ; ::_thesis: a = b hence a = b by A26; ::_thesis: verum end; suppose ( b <> {} & a <> {} & PosetMax b <> PosetMax a & [(PosetMax b),(PosetMax a)] in the InternalRel of R ) ; ::_thesis: a = b hence a = b by A26; ::_thesis: verum end; supposeA27: ( b <> {} & a <> {} & PosetMax b = PosetMax a & [(b \ {(PosetMax b)}),(a \ {(PosetMax a)})] in X ) ; ::_thesis: a = b reconsider a9 = a as finite set ; reconsider b9 = b as finite set ; set za = a9 \ {(PosetMax a)}; set zb = b9 \ {(PosetMax b)}; reconsider za = a9 \ {(PosetMax a)}, zb = b9 \ {(PosetMax b)} as Element of Fin the carrier of R by Th37; card za = n by A22, Lm1; then A28: za = zb by A21, A26, A27; now__::_thesis:_for_z_being_set_st_z_in_{(PosetMax_a)}_holds_ z_in_a let z be set ; ::_thesis: ( z in {(PosetMax a)} implies z in a ) assume z in {(PosetMax a)} ; ::_thesis: z in a then z = PosetMax a by TARSKI:def_1; hence z in a by A27, Def13; ::_thesis: verum end; then {(PosetMax a)} c= a by TARSKI:def_3; then A29: a = {(PosetMax a)} \/ za by XBOOLE_1:45; now__::_thesis:_for_z_being_set_st_z_in_{(PosetMax_b)}_holds_ z_in_b let z be set ; ::_thesis: ( z in {(PosetMax b)} implies z in b ) assume z in {(PosetMax b)} ; ::_thesis: z in b then z = PosetMax b by TARSKI:def_1; hence z in b by A27, Def13; ::_thesis: verum end; then {(PosetMax b)} c= b by TARSKI:def_3; hence a = b by A27, A28, A29, XBOOLE_1:45; ::_thesis: verum end; end; end; hence a = b ; ::_thesis: verum end; end; end; then A30: for n being Nat st S1[n] holds S1[n + 1] ; A31: for n being Nat holds S1[n] from NAT_1:sch_2(A20, A30); consider n being Nat such that A32: x,n are_equipotent by A14, CARD_1:43; card x9 = n by A32, CARD_1:def_2; hence x = y by A15, A16, A17, A31; ::_thesis: verum end; A33: now__::_thesis:_for_x,_y,_z_being_set_st_x_in_Fin_the_carrier_of_R_&_y_in_Fin_the_carrier_of_R_&_z_in_Fin_the_carrier_of_R_&_[x,y]_in_X_&_[y,z]_in_X_holds_ [x,z]_in_X let x, y, z be set ; ::_thesis: ( x in Fin the carrier of R & y in Fin the carrier of R & z in Fin the carrier of R & [x,y] in X & [y,z] in X implies [x,z] in X ) assume that A34: x in Fin the carrier of R and A35: y in Fin the carrier of R and A36: z in Fin the carrier of R and A37: [x,y] in X and A38: [y,z] in X ; ::_thesis: [x,z] in X defpred S1[ Nat] means for a, b, c being Element of Fin the carrier of R st card a = $1 & [a,b] in X & [b,c] in X holds [a,c] in X; now__::_thesis:_for_a,_b,_c_being_Element_of_Fin_the_carrier_of_R_st_card_a_=_0_&_[a,b]_in_X_&_[b,c]_in_X_holds_ [a,c]_in_X let a, b, c be Element of Fin the carrier of R; ::_thesis: ( card a = 0 & [a,b] in X & [b,c] in X implies [a,c] in X ) assume that A39: card a = 0 and [a,b] in X and [b,c] in X ; ::_thesis: [a,c] in X reconsider a9 = a as finite set ; a9 = {} by A39; hence [a,c] in X by Th35; ::_thesis: verum end; then A40: S1[ 0 ] ; now__::_thesis:_for_n_being_Nat_st_(_for_a,_b,_c_being_Element_of_Fin_the_carrier_of_R_st_card_a_=_n_&_[a,b]_in_X_&_[b,c]_in_X_holds_ [a,c]_in_X_)_holds_ for_a,_b,_c_being_Element_of_Fin_the_carrier_of_R_st_card_a_=_n_+_1_&_[a,b]_in_X_&_[b,c]_in_X_holds_ [a,c]_in_X let n be Nat; ::_thesis: ( ( for a, b, c being Element of Fin the carrier of R st card a = n & [a,b] in X & [b,c] in X holds [a,c] in X ) implies for a, b, c being Element of Fin the carrier of R st card a = n + 1 & [a,b] in X & [b,c] in X holds [b5,b7] in X ) assume A41: for a, b, c being Element of Fin the carrier of R st card a = n & [a,b] in X & [b,c] in X holds [a,c] in X ; ::_thesis: for a, b, c being Element of Fin the carrier of R st card a = n + 1 & [a,b] in X & [b,c] in X holds [b5,b7] in X let a, b, c be Element of Fin the carrier of R; ::_thesis: ( card a = n + 1 & [a,b] in X & [b,c] in X implies [b2,b4] in X ) assume that A42: card a = n + 1 and A43: [a,b] in X and A44: [b,c] in X ; ::_thesis: [b2,b4] in X percases ( a = {} or ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) or ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in union (rng (FinOrd-Approx R)) ) ) by A43, Th35; suppose a = {} ; ::_thesis: [b2,b4] in X hence [a,c] in X by Th35; ::_thesis: verum end; supposeA45: ( a <> {} & b <> {} & PosetMax a <> PosetMax b & [(PosetMax a),(PosetMax b)] in the InternalRel of R ) ; ::_thesis: [b2,b4] in X now__::_thesis:_[a,c]_in_X percases ( b = {} or ( b <> {} & c <> {} & PosetMax b <> PosetMax c & [(PosetMax b),(PosetMax c)] in the InternalRel of R ) or ( b <> {} & c <> {} & PosetMax b = PosetMax c & [(b \ {(PosetMax b)}),(c \ {(PosetMax c)})] in union (rng (FinOrd-Approx R)) ) ) by A44, Th35; suppose b = {} ; ::_thesis: [a,c] in X hence [a,c] in X by A45; ::_thesis: verum end; supposeA46: ( b <> {} & c <> {} & PosetMax b <> PosetMax c & [(PosetMax b),(PosetMax c)] in the InternalRel of R ) ; ::_thesis: [a,c] in X then A47: [(PosetMax a),(PosetMax c)] in the InternalRel of R by A45, ORDERS_1:5; now__::_thesis:_[a,c]_in_X percases ( PosetMax a <> PosetMax c or PosetMax a = PosetMax c ) ; suppose PosetMax a <> PosetMax c ; ::_thesis: [a,c] in X hence [a,c] in X by A45, A46, A47, Th35; ::_thesis: verum end; suppose PosetMax a = PosetMax c ; ::_thesis: [a,c] in X hence [a,c] in X by A45, A46, ORDERS_1:4; ::_thesis: verum end; end; end; hence [a,c] in X ; ::_thesis: verum end; suppose ( b <> {} & c <> {} & PosetMax b = PosetMax c & [(b \ {(PosetMax b)}),(c \ {(PosetMax c)})] in union (rng (FinOrd-Approx R)) ) ; ::_thesis: [a,c] in X hence [a,c] in X by A45, Th35; ::_thesis: verum end; end; end; hence [a,c] in X ; ::_thesis: verum end; supposeA48: ( a <> {} & b <> {} & PosetMax a = PosetMax b & [(a \ {(PosetMax a)}),(b \ {(PosetMax b)})] in union (rng (FinOrd-Approx R)) ) ; ::_thesis: [b2,b4] in X now__::_thesis:_[a,c]_in_X percases ( b = {} or ( b <> {} & c <> {} & PosetMax b <> PosetMax c & [(PosetMax b),(PosetMax c)] in the InternalRel of R ) or ( b <> {} & c <> {} & PosetMax b = PosetMax c & [(b \ {(PosetMax b)}),(c \ {(PosetMax c)})] in union (rng (FinOrd-Approx R)) ) ) by A44, Th35; suppose b = {} ; ::_thesis: [a,c] in X hence [a,c] in X by A48; ::_thesis: verum end; suppose ( b <> {} & c <> {} & PosetMax b <> PosetMax c & [(PosetMax b),(PosetMax c)] in the InternalRel of R ) ; ::_thesis: [a,c] in X hence [a,c] in X by A48, Th35; ::_thesis: verum end; supposeA49: ( b <> {} & c <> {} & PosetMax b = PosetMax c & [(b \ {(PosetMax b)}),(c \ {(PosetMax c)})] in union (rng (FinOrd-Approx R)) ) ; ::_thesis: [a,c] in X set z = a \ {(PosetMax a)}; reconsider z = a \ {(PosetMax a)} as Element of Fin the carrier of R by Th37; A50: card z = n by A42, Lm1; A51: c c= the carrier of R by FINSUB_1:def_5; reconsider c9 = c as finite set ; set zc = c9 \ {(PosetMax c)}; c9 \ {(PosetMax c)} c= the carrier of R by A51, XBOOLE_1:1; then reconsider zc = c9 \ {(PosetMax c)} as Element of Fin the carrier of R by FINSUB_1:def_5; A52: b c= the carrier of R by FINSUB_1:def_5; reconsider b9 = b as finite set ; set zb = b9 \ {(PosetMax b)}; b9 \ {(PosetMax b)} c= the carrier of R by A52, XBOOLE_1:1; then reconsider zb = b9 \ {(PosetMax b)} as Element of Fin the carrier of R by FINSUB_1:def_5; [z,zb] in union (rng (FinOrd-Approx R)) by A48; then [z,zc] in union (rng (FinOrd-Approx R)) by A41, A49, A50; hence [a,c] in X by A48, A49, Th35; ::_thesis: verum end; end; end; hence [a,c] in X ; ::_thesis: verum end; end; end; then A53: for n being Nat st S1[n] holds S1[n + 1] ; A54: for n being Nat holds S1[n] from NAT_1:sch_2(A40, A53); reconsider x9 = x as Element of Fin the carrier of R by A34; consider n being Nat such that A55: x,n are_equipotent by A34, CARD_1:43; card x9 = n by A55, CARD_1:def_2; hence [x,z] in X by A35, A36, A37, A38, A54; ::_thesis: verum end; A56: X is_reflexive_in Fin the carrier of R by A3, RELAT_2:def_1; A57: X is_antisymmetric_in Fin the carrier of R by A13, RELAT_2:def_4; A58: X is_transitive_in Fin the carrier of R by A33, RELAT_2:def_8; reconsider R = union (rng (FinOrd-Approx R)) as Relation of (Fin the carrier of R) by A3; A59: dom R = Fin the carrier of R by A56, ORDERS_1:13; field R = Fin the carrier of R by A56, ORDERS_1:13; hence union (rng (FinOrd-Approx R)) is Order of (Fin the carrier of R) by A56, A57, A58, A59, PARTFUN1:def_2, RELAT_2:def_9, RELAT_2:def_12, RELAT_2:def_16; ::_thesis: verum end; definition let R be non empty connected Poset; func FinOrd R -> Order of (Fin the carrier of R) equals :: BAGORDER:def 15 union (rng (FinOrd-Approx R)); coherence union (rng (FinOrd-Approx R)) is Order of (Fin the carrier of R) by Th38; end; :: deftheorem defines FinOrd BAGORDER:def_15_:_ for R being non empty connected Poset holds FinOrd R = union (rng (FinOrd-Approx R)); definition let R be non empty connected Poset; func FinPoset R -> Poset equals :: BAGORDER:def 16 RelStr(# (Fin the carrier of R),(FinOrd R) #); correctness coherence RelStr(# (Fin the carrier of R),(FinOrd R) #) is Poset; ; end; :: deftheorem defines FinPoset BAGORDER:def_16_:_ for R being non empty connected Poset holds FinPoset R = RelStr(# (Fin the carrier of R),(FinOrd R) #); registration let R be non empty connected Poset; cluster FinPoset R -> non empty ; correctness coherence not FinPoset R is empty ; ; end; theorem Th39: :: BAGORDER:39 for R being non empty connected Poset for a, b being Element of (FinPoset R) holds ( [a,b] in the InternalRel of (FinPoset R) iff ex x, y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) ) proof let R be non empty connected Poset; ::_thesis: for a, b being Element of (FinPoset R) holds ( [a,b] in the InternalRel of (FinPoset R) iff ex x, y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) ) let a, b be Element of (FinPoset R); ::_thesis: ( [a,b] in the InternalRel of (FinPoset R) iff ex x, y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) ) set CR = the carrier of R; reconsider x = a, y = b as Element of Fin the carrier of R ; hereby ::_thesis: ( ex x, y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) implies [a,b] in the InternalRel of (FinPoset R) ) assume A1: [a,b] in the InternalRel of (FinPoset R) ; ::_thesis: ex x, y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) take x = x; ::_thesis: ex y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) take y = y; ::_thesis: ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) thus ( a = x & b = y ) ; ::_thesis: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) thus ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) by A1, Th35; ::_thesis: verum end; assume ex x, y being Element of Fin the carrier of R st ( a = x & b = y & ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) ) ; ::_thesis: [a,b] in the InternalRel of (FinPoset R) hence [a,b] in the InternalRel of (FinPoset R) by Th35; ::_thesis: verum end; registration let R be non empty connected Poset; cluster FinPoset R -> connected ; correctness coherence FinPoset R is connected ; proof set IR = the InternalRel of R; set CR = the carrier of R; set FIR = FinOrd R; set FCR = Fin the carrier of R; A1: FinPoset R = RelStr(# (Fin the carrier of R),(FinOrd R) #) ; now__::_thesis:_for_x,_y_being_Element_of_(FinPoset_R)_holds_ (_x_<=_y_or_y_<=_x_) let x, y be Element of (FinPoset R); ::_thesis: ( x <= y or y <= x ) reconsider x9 = x, y9 = y as Element of Fin the carrier of R ; defpred S1[ Nat] means for x, y being Element of Fin the carrier of R holds ( not card x = R or [x,y] in FinOrd R or [y,x] in FinOrd R ); now__::_thesis:_for_a,_b_being_Element_of_Fin_the_carrier_of_R_holds_ (_not_card_a_=_0_or_[a,b]_in_FinOrd_R_or_[a,b]_in_FinOrd_R_) let a, b be Element of Fin the carrier of R; ::_thesis: ( not card a = 0 or [a,b] in FinOrd R or [a,b] in FinOrd R ) assume card a = 0 ; ::_thesis: ( [a,b] in FinOrd R or [a,b] in FinOrd R ) then a = {} ; hence ( [a,b] in FinOrd R or [a,b] in FinOrd R ) by A1, Th39; ::_thesis: verum end; then A2: S1[ 0 ] ; now__::_thesis:_for_n_being_Nat_st_(_for_x,_y_being_Element_of_Fin_the_carrier_of_R_holds_ (_not_card_x_=_n_or_[x,y]_in_FinOrd_R_or_[y,x]_in_FinOrd_R_)_)_holds_ for_a,_b_being_Element_of_Fin_the_carrier_of_R_holds_ (_not_card_a_=_n_+_1_or_[a,b]_in_FinOrd_R_or_[b,a]_in_FinOrd_R_) let n be Nat; ::_thesis: ( ( for x, y being Element of Fin the carrier of R holds ( not card x = n or [x,y] in FinOrd R or [y,x] in FinOrd R ) ) implies for a, b being Element of Fin the carrier of R holds ( not card a = n + 1 or [b4,b5] in FinOrd R or [b5,b4] in FinOrd R ) ) assume A3: for x, y being Element of Fin the carrier of R holds ( not card x = n or [x,y] in FinOrd R or [y,x] in FinOrd R ) ; ::_thesis: for a, b being Element of Fin the carrier of R holds ( not card a = n + 1 or [b4,b5] in FinOrd R or [b5,b4] in FinOrd R ) let a, b be Element of Fin the carrier of R; ::_thesis: ( not card a = n + 1 or [b2,b3] in FinOrd R or [b3,b2] in FinOrd R ) assume A4: card a = n + 1 ; ::_thesis: ( [b2,b3] in FinOrd R or [b3,b2] in FinOrd R ) percases ( a = {} or a <> {} ) ; suppose a = {} ; ::_thesis: ( [b2,b3] in FinOrd R or [b3,b2] in FinOrd R ) hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) by A1, Th39; ::_thesis: verum end; supposeA5: a <> {} ; ::_thesis: ( [b2,b3] in FinOrd R or [b3,b2] in FinOrd R ) now__::_thesis:_(_[a,b]_in_FinOrd_R_or_[b,a]_in_FinOrd_R_) percases ( b = {} or b <> {} ) ; suppose b = {} ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) by A1, Th39; ::_thesis: verum end; supposeA6: b <> {} ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) now__::_thesis:_(_[a,b]_in_FinOrd_R_or_[b,a]_in_FinOrd_R_) percases ( PosetMax a <> PosetMax b or PosetMax a = PosetMax b ) ; supposeA7: PosetMax a <> PosetMax b ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) now__::_thesis:_(_[a,b]_in_FinOrd_R_or_[b,a]_in_FinOrd_R_) percases ( PosetMax a <= PosetMax b or PosetMax b <= PosetMax a ) by WAYBEL_0:def_29; suppose PosetMax a <= PosetMax b ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) then [(PosetMax a),(PosetMax b)] in the InternalRel of R by ORDERS_2:def_5; hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) by A1, A5, A6, A7, Th39; ::_thesis: verum end; suppose PosetMax b <= PosetMax a ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) then [(PosetMax b),(PosetMax a)] in the InternalRel of R by ORDERS_2:def_5; hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) by A1, A5, A6, A7, Th39; ::_thesis: verum end; end; end; hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) ; ::_thesis: verum end; supposeA8: PosetMax a = PosetMax b ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) set ax = a \ {(PosetMax a)}; set bx = b \ {(PosetMax b)}; A9: card (a \ {(PosetMax a)}) = n by A4, Lm1; reconsider ax = a \ {(PosetMax a)}, bx = b \ {(PosetMax b)} as Element of Fin the carrier of R by Th37; now__::_thesis:_(_[a,b]_in_FinOrd_R_or_[b,a]_in_FinOrd_R_) percases ( [ax,bx] in FinOrd R or [bx,ax] in FinOrd R ) by A3, A9; suppose [ax,bx] in FinOrd R ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) by A1, A5, A6, A8, Th39; ::_thesis: verum end; suppose [bx,ax] in FinOrd R ; ::_thesis: ( [a,b] in FinOrd R or [b,a] in FinOrd R ) hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) by A1, A5, A6, A8, Th39; ::_thesis: verum end; end; end; hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) ; ::_thesis: verum end; end; end; hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) ; ::_thesis: verum end; end; end; hence ( [a,b] in FinOrd R or [b,a] in FinOrd R ) ; ::_thesis: verum end; end; end; then A10: for n being Nat st S1[n] holds S1[n + 1] ; A11: for n being Nat holds S1[n] from NAT_1:sch_2(A2, A10); consider n being Nat such that A12: x,n are_equipotent by CARD_1:43; card x9 = n by A12, CARD_1:def_2; then ( [x9,y9] in FinOrd R or [y9,x9] in FinOrd R ) by A11; hence ( x <= y or y <= x ) by ORDERS_2:def_5; ::_thesis: verum end; hence FinPoset R is connected by WAYBEL_0:def_29; ::_thesis: verum end; end; definition let R be non empty connected RelStr ; let C be non empty set ; assume that B1: R is well_founded and B2: C c= the carrier of R ; func MinElement (C,R) -> Element of R means :Def17: :: BAGORDER:def 17 ( it in C & it is_minimal_wrt C, the InternalRel of R ); existence ex b1 being Element of R st ( b1 in C & b1 is_minimal_wrt C, the InternalRel of R ) proof set IR = the InternalRel of R; set CR = the carrier of R; the InternalRel of R is_well_founded_in the carrier of R by B1, WELLFND1:def_2; then consider a0 being set such that A1: a0 in C and A2: the InternalRel of R -Seg a0 misses C by B2, WELLORD1:def_3; reconsider a0 = a0 as Element of R by B2, A1; take a0 ; ::_thesis: ( a0 in C & a0 is_minimal_wrt C, the InternalRel of R ) thus a0 in C by A1; ::_thesis: a0 is_minimal_wrt C, the InternalRel of R thus a0 is_minimal_wrt C, the InternalRel of R by A1, A2, DICKSON:6; ::_thesis: verum end; uniqueness for b1, b2 being Element of R st b1 in C & b1 is_minimal_wrt C, the InternalRel of R & b2 in C & b2 is_minimal_wrt C, the InternalRel of R holds b1 = b2 proof let IT1, IT2 be Element of R; ::_thesis: ( IT1 in C & IT1 is_minimal_wrt C, the InternalRel of R & IT2 in C & IT2 is_minimal_wrt C, the InternalRel of R implies IT1 = IT2 ) assume that A3: IT1 in C and A4: IT1 is_minimal_wrt C, the InternalRel of R and A5: IT2 in C and A6: IT2 is_minimal_wrt C, the InternalRel of R ; ::_thesis: IT1 = IT2 set IR = the InternalRel of R; assume A7: IT1 <> IT2 ; ::_thesis: contradiction percases ( IT1 <= IT2 or IT2 <= IT1 ) by WAYBEL_0:def_29; suppose IT1 <= IT2 ; ::_thesis: contradiction then [IT1,IT2] in the InternalRel of R by ORDERS_2:def_5; then IT1 in the InternalRel of R -Seg IT2 by A7, WELLORD1:1; then IT1 in ( the InternalRel of R -Seg IT2) /\ C by A3, XBOOLE_0:def_4; then the InternalRel of R -Seg IT2 meets C by XBOOLE_0:def_7; hence contradiction by A6, DICKSON:6; ::_thesis: verum end; suppose IT2 <= IT1 ; ::_thesis: contradiction then [IT2,IT1] in the InternalRel of R by ORDERS_2:def_5; then IT2 in the InternalRel of R -Seg IT1 by A7, WELLORD1:1; then IT2 in ( the InternalRel of R -Seg IT1) /\ C by A5, XBOOLE_0:def_4; then the InternalRel of R -Seg IT1 meets C by XBOOLE_0:def_7; hence contradiction by A4, DICKSON:6; ::_thesis: verum end; end; end; end; :: deftheorem Def17 defines MinElement BAGORDER:def_17_:_ for R being non empty connected RelStr for C being non empty set st R is well_founded & C c= the carrier of R holds for b3 being Element of R holds ( b3 = MinElement (C,R) iff ( b3 in C & b3 is_minimal_wrt C, the InternalRel of R ) ); theorem Th40: :: BAGORDER:40 for R being non empty RelStr for s being sequence of R for j being Nat st s is descending holds s ^\ j is descending proof let R be non empty RelStr ; ::_thesis: for s being sequence of R for j being Nat st s is descending holds s ^\ j is descending let s1 be sequence of R; ::_thesis: for j being Nat st s1 is descending holds s1 ^\ j is descending let j be Nat; ::_thesis: ( s1 is descending implies s1 ^\ j is descending ) assume A1: s1 is descending ; ::_thesis: s1 ^\ j is descending set s2 = s1 ^\ j; set IR = the InternalRel of R; now__::_thesis:_for_n_being_Nat_holds_ (_(s1_^\_j)_._(n_+_1)_<>_(s1_^\_j)_._n_&_[((s1_^\_j)_._(n_+_1)),((s1_^\_j)_._n)]_in_the_InternalRel_of_R_) let n be Nat; ::_thesis: ( (s1 ^\ j) . (n + 1) <> (s1 ^\ j) . n & [((s1 ^\ j) . (n + 1)),((s1 ^\ j) . n)] in the InternalRel of R ) set nj = n + j; A2: (s1 ^\ j) . n = s1 . (n + j) by NAT_1:def_3; A3: (s1 ^\ j) . (n + 1) = s1 . ((n + 1) + j) by NAT_1:def_3 .= s1 . ((n + j) + 1) ; hence (s1 ^\ j) . (n + 1) <> (s1 ^\ j) . n by A1, A2, WELLFND1:def_6; ::_thesis: [((s1 ^\ j) . (n + 1)),((s1 ^\ j) . n)] in the InternalRel of R thus [((s1 ^\ j) . (n + 1)),((s1 ^\ j) . n)] in the InternalRel of R by A1, A2, A3, WELLFND1:def_6; ::_thesis: verum end; hence s1 ^\ j is descending by WELLFND1:def_6; ::_thesis: verum end; theorem :: BAGORDER:41 for R being non empty connected Poset st R is well_founded holds FinPoset R is well_founded proof let R be non empty connected Poset; ::_thesis: ( R is well_founded implies FinPoset R is well_founded ) assume A1: R is well_founded ; ::_thesis: FinPoset R is well_founded set IR = the InternalRel of R; set CR = the carrier of R; set FIR = FinOrd R; set FCR = Fin the carrier of R; assume not FinPoset R is well_founded ; ::_thesis: contradiction then consider A being sequence of (FinPoset R) such that A2: A is descending by WELLFND1:14; set zz = { z where z is sequence of (FinPoset R) : z is descending } ; A in { z where z is sequence of (FinPoset R) : z is descending } by A2; then reconsider zz = { z where z is sequence of (FinPoset R) : z is descending } as non empty set ; set Z = [: the carrier of R,zz:]; defpred S1[ Nat, set , set ] means ex Sn2 being sequence of (FinPoset R) ex Smax being Function of NAT, the carrier of R ex an being Element of R ex ix being Nat ex bnt, bn being sequence of (FinPoset R) st ( Sn2 = $2 `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & an = MinElement ((rng Smax),R) & ix = Smax mindex an & bnt = Sn2 ^\ ix & ( for i being Nat holds bn . i = (bnt . i) \ {an} ) & ( for i being Nat st ix <= i holds Smax . i = an ) & $3 = [an,bn] ); A3: for n being Element of NAT for Sn being Element of [: the carrier of R,zz:] ex Sn1 being Element of [: the carrier of R,zz:] st S1[n,Sn,Sn1] proof let n be Element of NAT ; ::_thesis: for Sn being Element of [: the carrier of R,zz:] ex Sn1 being Element of [: the carrier of R,zz:] st S1[n,Sn,Sn1] let Sn be Element of [: the carrier of R,zz:]; ::_thesis: ex Sn1 being Element of [: the carrier of R,zz:] st S1[n,Sn,Sn1] set Sn2 = Sn `2 ; Sn `2 in zz ; then A4: ex z being sequence of (FinPoset R) st ( z = Sn `2 & z is descending ) ; then reconsider Sn2 = Sn `2 as sequence of (FinPoset R) ; A5: now__::_thesis:_for_i_being_Nat_holds_not_Sn2_._i_=_{} let i be Nat; ::_thesis: not Sn2 . i = {} assume A6: Sn2 . i = {} ; ::_thesis: contradiction A7: Sn2 . (i + 1) <> Sn2 . i by A4, WELLFND1:def_6; [(Sn2 . (i + 1)),(Sn2 . i)] in FinOrd R by A4, WELLFND1:def_6; hence contradiction by A6, A7, Th36; ::_thesis: verum end; defpred S2[ Nat, set ] means ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . $1 & Sn2i <> {} & $2 = PosetMax Sn2i ); A8: for i being Element of NAT ex y being Element of the carrier of R st S2[i,y] proof let i be Element of NAT ; ::_thesis: ex y being Element of the carrier of R st S2[i,y] set Sn2i = Sn2 . i; reconsider Sn2i = Sn2 . i as Element of Fin the carrier of R ; set y = PosetMax Sn2i; take PosetMax Sn2i ; ::_thesis: S2[i, PosetMax Sn2i] take Sn2i ; ::_thesis: ( Sn2i = Sn2 . i & Sn2i <> {} & PosetMax Sn2i = PosetMax Sn2i ) thus Sn2i = Sn2 . i ; ::_thesis: ( Sn2i <> {} & PosetMax Sn2i = PosetMax Sn2i ) thus Sn2i <> {} by A5; ::_thesis: PosetMax Sn2i = PosetMax Sn2i thus PosetMax Sn2i = PosetMax Sn2i ; ::_thesis: verum end; consider Smax being Function of NAT, the carrier of R such that A9: for i being Element of NAT holds S2[i,Smax . i] from FUNCT_2:sch_3(A8); set an = MinElement ((rng Smax),R); set ix = Smax mindex (MinElement ((rng Smax),R)); set bnt = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))); defpred S3[ set , set ] means ex bni being Element of Fin the carrier of R st ( bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . $1) \ {(MinElement ((rng Smax),R))} & $2 = bni ); now__::_thesis:_for_i_being_Nat_ex_y,_bni_being_Element_of_Fin_the_carrier_of_R_st_ (_bni_=_((Sn2_^\_(Smax_mindex_(MinElement_((rng_Smax),R))))_._i)_\_{(MinElement_((rng_Smax),R))}_&_y_=_bni_) let i be Nat; ::_thesis: ex y, bni being Element of Fin the carrier of R st ( bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} & y = bni ) set bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))}; reconsider k = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i as finite Subset of the carrier of R by FINSUB_1:def_5; k \ {(MinElement ((rng Smax),R))} in Fin the carrier of R by FINSUB_1:def_5; then reconsider bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} as Element of Fin the carrier of R ; set y = bni; take y = bni; ::_thesis: ex bni being Element of Fin the carrier of R st ( bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} & y = bni ) take bni = bni; ::_thesis: ( bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} & y = bni ) thus bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ; ::_thesis: y = bni thus y = bni ; ::_thesis: verum end; then A10: for i being Element of NAT ex y being Element of Fin the carrier of R st S3[i,y] ; defpred S4[ Nat] means Smax . ((Smax mindex (MinElement ((rng Smax),R))) + $1) = MinElement ((rng Smax),R); A11: dom Smax = NAT by FUNCT_2:def_1; A12: rng Smax c= the carrier of R by RELAT_1:def_19; then A13: MinElement ((rng Smax),R) in rng Smax by A1, Def17; A14: MinElement ((rng Smax),R) is_minimal_wrt rng Smax, the InternalRel of R by A1, A12, Def17; A15: S4[ 0 ] by A11, A13, DICKSON:def_11; A16: now__::_thesis:_for_k_being_Nat_st_S4[k]_holds_ S4[k_+_1] let k be Nat; ::_thesis: ( S4[k] implies S4[b1 + 1] ) assume A17: S4[k] ; ::_thesis: S4[b1 + 1] set ixk = (Smax mindex (MinElement ((rng Smax),R))) + k; set ixk1 = (Smax mindex (MinElement ((rng Smax),R))) + (k + 1); set ixk19 = ((Smax mindex (MinElement ((rng Smax),R))) + k) + 1; consider Sn2ixk being Element of Fin the carrier of R such that A18: Sn2ixk = Sn2 . ((Smax mindex (MinElement ((rng Smax),R))) + k) and Sn2ixk <> {} and A19: Smax . ((Smax mindex (MinElement ((rng Smax),R))) + k) = PosetMax Sn2ixk by A9; consider Sn2ixk1 being Element of Fin the carrier of R such that A20: Sn2ixk1 = Sn2 . ((Smax mindex (MinElement ((rng Smax),R))) + (k + 1)) and A21: Sn2ixk1 <> {} and A22: Smax . ((Smax mindex (MinElement ((rng Smax),R))) + (k + 1)) = PosetMax Sn2ixk1 by A9; reconsider Sn2ixk9 = Sn2ixk, Sn2ixk19 = Sn2ixk1 as Element of (FinPoset R) ; (Smax mindex (MinElement ((rng Smax),R))) + (k + 1) = ((Smax mindex (MinElement ((rng Smax),R))) + k) + 1 ; then [Sn2ixk19,Sn2ixk9] in FinOrd R by A4, A18, A20, WELLFND1:def_6; then consider x, y being Element of Fin the carrier of R such that A23: Sn2ixk1 = x and A24: Sn2ixk = y and A25: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) by Th39; percases ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) by A25; suppose x = {} ; ::_thesis: S4[b1 + 1] hence S4[k + 1] by A21, A23; ::_thesis: verum end; supposeA26: ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ; ::_thesis: S4[b1 + 1] Smax . ((Smax mindex (MinElement ((rng Smax),R))) + (k + 1)) in rng Smax by A11, FUNCT_1:def_3; hence S4[k + 1] by A14, A17, A19, A22, A23, A24, A26, WAYBEL_4:def_25; ::_thesis: verum end; suppose ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ; ::_thesis: S4[b1 + 1] hence S4[k + 1] by A17, A19, A22, A23, A24; ::_thesis: verum end; end; end; A27: for k being Nat holds S4[k] from NAT_1:sch_2(A15, A16); A28: now__::_thesis:_for_i_being_Nat_st_Smax_mindex_(MinElement_((rng_Smax),R))_<=_i_holds_ Smax_._i_=_MinElement_((rng_Smax),R) let i be Nat; ::_thesis: ( Smax mindex (MinElement ((rng Smax),R)) <= i implies Smax . i = MinElement ((rng Smax),R) ) assume Smax mindex (MinElement ((rng Smax),R)) <= i ; ::_thesis: Smax . i = MinElement ((rng Smax),R) then consider k being Nat such that A29: i = (Smax mindex (MinElement ((rng Smax),R))) + k by NAT_1:10; reconsider k = k as Nat ; i = (Smax mindex (MinElement ((rng Smax),R))) + k by A29; hence Smax . i = MinElement ((rng Smax),R) by A27; ::_thesis: verum end; A30: now__::_thesis:_for_i_being_Nat_st_Smax_mindex_(MinElement_((rng_Smax),R))_<=_i_holds_ ex_Sn2i_being_Element_of_Fin_the_carrier_of_R_st_ (_Sn2i_=_Sn2_._i_&_PosetMax_Sn2i_=_MinElement_((rng_Smax),R)_) let i be Nat; ::_thesis: ( Smax mindex (MinElement ((rng Smax),R)) <= i implies ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & PosetMax Sn2i = MinElement ((rng Smax),R) ) ) assume A31: Smax mindex (MinElement ((rng Smax),R)) <= i ; ::_thesis: ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & PosetMax Sn2i = MinElement ((rng Smax),R) ) reconsider i0 = i as Element of NAT by ORDINAL1:def_12; consider Sn2i being Element of Fin the carrier of R such that A32: Sn2i = Sn2 . i and Sn2i <> {} and A33: Smax . i0 = PosetMax Sn2i by A9; take Sn2i = Sn2i; ::_thesis: ( Sn2i = Sn2 . i & PosetMax Sn2i = MinElement ((rng Smax),R) ) thus Sn2i = Sn2 . i by A32; ::_thesis: PosetMax Sn2i = MinElement ((rng Smax),R) thus PosetMax Sn2i = MinElement ((rng Smax),R) by A28, A31, A33; ::_thesis: verum end; A34: now__::_thesis:_for_i_being_Nat_ex_bnti_being_Element_of_Fin_the_carrier_of_R_st_ (_bnti_=_(Sn2_^\_(Smax_mindex_(MinElement_((rng_Smax),R))))_._i_&_PosetMax_bnti_=_MinElement_((rng_Smax),R)_) let i be Nat; ::_thesis: ex bnti being Element of Fin the carrier of R st ( bnti = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i & PosetMax bnti = MinElement ((rng Smax),R) ) set bnti = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i; reconsider bnti = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i as Element of Fin the carrier of R ; take bnti = bnti; ::_thesis: ( bnti = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i & PosetMax bnti = MinElement ((rng Smax),R) ) thus bnti = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i ; ::_thesis: PosetMax bnti = MinElement ((rng Smax),R) set iix = i + (Smax mindex (MinElement ((rng Smax),R))); ex Sn2iix being Element of Fin the carrier of R st ( Sn2iix = Sn2 . (i + (Smax mindex (MinElement ((rng Smax),R)))) & PosetMax Sn2iix = MinElement ((rng Smax),R) ) by A30, NAT_1:11; hence PosetMax bnti = MinElement ((rng Smax),R) by NAT_1:def_3; ::_thesis: verum end; consider bn being Function of NAT,(Fin the carrier of R) such that A35: for i being Element of NAT holds S3[i,bn . i] from FUNCT_2:sch_3(A10); reconsider bn = bn as sequence of (FinPoset R) ; set Sn1 = [(MinElement ((rng Smax),R)),bn]; A36: Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) is descending by A4, Th40; now__::_thesis:_for_i_being_Nat_holds_ (_bn_._(i_+_1)_<>_bn_._i_&_[(bn_._(i_+_1)),(bn_._i)]_in_FinOrd_R_) let i be Nat; ::_thesis: ( bn . (i + 1) <> bn . i & [(bn . (b1 + 1)),(bn . b1)] in FinOrd R ) reconsider i0 = i as Element of NAT by ORDINAL1:def_12; A37: ex bni being Element of Fin the carrier of R st ( bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i0) \ {(MinElement ((rng Smax),R))} & bn . i0 = bni ) by A35; A38: ex bni1 being Element of Fin the carrier of R st ( bni1 = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . (i + 1)) \ {(MinElement ((rng Smax),R))} & bn . (i + 1) = bni1 ) by A35; reconsider bnti = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i, bnti1 = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . (i + 1) as Element of (FinPoset R) ; reconsider bnti9 = bnti, bnti19 = bnti1 as Element of Fin the carrier of R ; A39: bnti1 <> bnti by A36, WELLFND1:def_6; [bnti1,bnti] in FinOrd R by A36, WELLFND1:def_6; then consider x, y being Element of Fin the carrier of R such that A40: bnti1 = x and A41: bnti = y and A42: ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) by Th39; A43: now__::_thesis:_for_i_being_Nat_holds_(Sn2_^\_(Smax_mindex_(MinElement_((rng_Smax),R))))_._i_<>_{} let i be Nat; ::_thesis: (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i <> {} (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i = Sn2 . (i + (Smax mindex (MinElement ((rng Smax),R)))) by NAT_1:def_3; hence (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i <> {} by A5; ::_thesis: verum end; A44: now__::_thesis:_(_PosetMax_bnti9_=_MinElement_((rng_Smax),R)_&_PosetMax_bnti19_=_MinElement_((rng_Smax),R)_) A45: ex bnti99 being Element of Fin the carrier of R st ( bnti99 = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i & PosetMax bnti99 = MinElement ((rng Smax),R) ) by A34; ex bnti199 being Element of Fin the carrier of R st ( bnti199 = (Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . (i + 1) & PosetMax bnti199 = MinElement ((rng Smax),R) ) by A34; hence ( PosetMax bnti9 = MinElement ((rng Smax),R) & PosetMax bnti19 = MinElement ((rng Smax),R) ) by A45; ::_thesis: verum end; A46: bnti9 <> {} by A43; A47: bnti19 <> {} by A43; A48: MinElement ((rng Smax),R) in bnti by A44, A46, Def13; MinElement ((rng Smax),R) in bnti1 by A44, A47, Def13; hence bn . (i + 1) <> bn . i by A37, A38, A39, A48, Th1; ::_thesis: [(bn . (b1 + 1)),(bn . b1)] in FinOrd R percases ( x = {} or ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) or ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ) by A42; suppose x = {} ; ::_thesis: [(bn . (b1 + 1)),(bn . b1)] in FinOrd R hence [(bn . (i + 1)),(bn . i)] in FinOrd R by A40, A43; ::_thesis: verum end; suppose ( x <> {} & y <> {} & PosetMax x <> PosetMax y & [(PosetMax x),(PosetMax y)] in the InternalRel of R ) ; ::_thesis: [(bn . (b1 + 1)),(bn . b1)] in FinOrd R hence [(bn . (i + 1)),(bn . i)] in FinOrd R by A40, A41, A44; ::_thesis: verum end; suppose ( x <> {} & y <> {} & PosetMax x = PosetMax y & [(x \ {(PosetMax x)}),(y \ {(PosetMax y)})] in FinOrd R ) ; ::_thesis: [(bn . (b1 + 1)),(bn . b1)] in FinOrd R hence [(bn . (i + 1)),(bn . i)] in FinOrd R by A37, A38, A40, A41, A44; ::_thesis: verum end; end; end; then bn is descending by WELLFND1:def_6; then bn in zz ; then reconsider Sn1 = [(MinElement ((rng Smax),R)),bn] as Element of [: the carrier of R,zz:] by ZFMISC_1:def_2; take Sn1 ; ::_thesis: S1[n,Sn,Sn1] take Sn2 ; ::_thesis: ex Smax being Function of NAT, the carrier of R ex an being Element of R ex ix being Nat ex bnt, bn being sequence of (FinPoset R) st ( Sn2 = Sn `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & an = MinElement ((rng Smax),R) & ix = Smax mindex an & bnt = Sn2 ^\ ix & ( for i being Nat holds bn . i = (bnt . i) \ {an} ) & ( for i being Nat st ix <= i holds Smax . i = an ) & Sn1 = [an,bn] ) take Smax ; ::_thesis: ex an being Element of R ex ix being Nat ex bnt, bn being sequence of (FinPoset R) st ( Sn2 = Sn `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & an = MinElement ((rng Smax),R) & ix = Smax mindex an & bnt = Sn2 ^\ ix & ( for i being Nat holds bn . i = (bnt . i) \ {an} ) & ( for i being Nat st ix <= i holds Smax . i = an ) & Sn1 = [an,bn] ) take MinElement ((rng Smax),R) ; ::_thesis: ex ix being Nat ex bnt, bn being sequence of (FinPoset R) st ( Sn2 = Sn `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & MinElement ((rng Smax),R) = MinElement ((rng Smax),R) & ix = Smax mindex (MinElement ((rng Smax),R)) & bnt = Sn2 ^\ ix & ( for i being Nat holds bn . i = (bnt . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st ix <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) take Smax mindex (MinElement ((rng Smax),R)) ; ::_thesis: ex bnt, bn being sequence of (FinPoset R) st ( Sn2 = Sn `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & MinElement ((rng Smax),R) = MinElement ((rng Smax),R) & Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) & bnt = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = (bnt . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) take Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) ; ::_thesis: ex bn being sequence of (FinPoset R) st ( Sn2 = Sn `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & MinElement ((rng Smax),R) = MinElement ((rng Smax),R) & Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) & Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) take bn ; ::_thesis: ( Sn2 = Sn `2 & ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & MinElement ((rng Smax),R) = MinElement ((rng Smax),R) & Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) & Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) thus Sn2 = Sn `2 ; ::_thesis: ( ( for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ) & MinElement ((rng Smax),R) = MinElement ((rng Smax),R) & Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) & Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) thus for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ::_thesis: ( MinElement ((rng Smax),R) = MinElement ((rng Smax),R) & Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) & Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) proof let i be Nat; ::_thesis: ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) reconsider i0 = i as Element of NAT by ORDINAL1:def_12; ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i0 = PosetMax Sn2i ) by A9; hence ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) ; ::_thesis: verum end; thus MinElement ((rng Smax),R) = MinElement ((rng Smax),R) ; ::_thesis: ( Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) & Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) thus Smax mindex (MinElement ((rng Smax),R)) = Smax mindex (MinElement ((rng Smax),R)) ; ::_thesis: ( Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) & ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) thus Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) = Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R))) ; ::_thesis: ( ( for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ) & ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) now__::_thesis:_for_i_being_Nat_holds_bn_._i_=_((Sn2_^\_(Smax_mindex_(MinElement_((rng_Smax),R))))_._i)_\_{(MinElement_((rng_Smax),R))} let i be Nat; ::_thesis: bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} reconsider i0 = i as Element of NAT by ORDINAL1:def_12; ex bni being Element of Fin the carrier of R st ( bni = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i0) \ {(MinElement ((rng Smax),R))} & bn . i0 = bni ) by A35; hence bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ; ::_thesis: verum end; hence for i being Nat holds bn . i = ((Sn2 ^\ (Smax mindex (MinElement ((rng Smax),R)))) . i) \ {(MinElement ((rng Smax),R))} ; ::_thesis: ( ( for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) ) & Sn1 = [(MinElement ((rng Smax),R)),bn] ) thus for i being Nat st Smax mindex (MinElement ((rng Smax),R)) <= i holds Smax . i = MinElement ((rng Smax),R) by A28; ::_thesis: Sn1 = [(MinElement ((rng Smax),R)),bn] thus Sn1 = [(MinElement ((rng Smax),R)),bn] ; ::_thesis: verum end; set aStart = the Element of R; set SS = [ the Element of R,A]; A in zz by A2; then reconsider SS = [ the Element of R,A] as Element of [: the carrier of R,zz:] by ZFMISC_1:def_2; consider S01 being Element of [: the carrier of R,zz:], S02 being sequence of (FinPoset R), S0max being Function of NAT, the carrier of R, a0 being Element of R, i0 being Nat, b0t, b0 being sequence of (FinPoset R) such that S02 = SS `2 and A49: for i being Nat ex S02i being Element of Fin the carrier of R st ( S02i = S02 . i & S02i <> {} & S0max . i = PosetMax S02i ) and a0 = MinElement ((rng S0max),R) and i0 = S0max mindex a0 and A50: b0t = S02 ^\ i0 and A51: for i being Nat holds b0 . i = (b0t . i) \ {a0} and A52: for i being Nat st i0 <= i holds S0max . i = a0 and A53: S01 = [a0,b0] by A3; consider S being Function of NAT,[: the carrier of R,zz:] such that A54: S . 0 = S01 and A55: for n being Element of NAT holds S1[n,S . n,S . (n + 1)] from RECDEF_1:sch_2(A3); A56: for n being Nat holds S1[n,S . n,S . (n + 1)] proof let n be Nat; ::_thesis: S1[n,S . n,S . (n + 1)] reconsider n0 = n as Element of NAT by ORDINAL1:def_12; S1[n0,S . n0,S . (n0 + 1)] by A55; hence S1[n,S . n,S . (n + 1)] ; ::_thesis: verum end; deffunc H1( set ) -> set = (S . $1) `1 ; A57: 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 Nat ; reconsider Sx = S . x9 as Element of [: the carrier of R,zz:] ; Sx `1 in the carrier of R ; hence H1(x) in the carrier of R ; ::_thesis: verum end; consider a being Function of NAT, the carrier of R such that A58: for x being set st x in NAT holds a . x = H1(x) from FUNCT_2:sch_2(A57); reconsider a = a as sequence of R ; defpred S2[ Nat] means for i being Nat ex b being sequence of (FinPoset R) st ( b = (S . $1) `2 & ( for x being set st x in b . i holds ( x <> (S . $1) `1 & [x,((S . $1) `1)] in the InternalRel of R ) ) ); A59: S2[ 0 ] proof let i be Nat; ::_thesis: ex b being sequence of (FinPoset R) st ( b = (S . 0) `2 & ( for x being set st x in b . i holds ( x <> (S . 0) `1 & [x,((S . 0) `1)] in the InternalRel of R ) ) ) set b = (S . 0) `2 ; (S . 0) `2 in zz ; then ex z being sequence of (FinPoset R) st ( z = (S . 0) `2 & z is descending ) ; then reconsider b = (S . 0) `2 as sequence of (FinPoset R) ; take b ; ::_thesis: ( b = (S . 0) `2 & ( for x being set st x in b . i holds ( x <> (S . 0) `1 & [x,((S . 0) `1)] in the InternalRel of R ) ) ) thus b = (S . 0) `2 ; ::_thesis: for x being set st x in b . i holds ( x <> (S . 0) `1 & [x,((S . 0) `1)] in the InternalRel of R ) let x be set ; ::_thesis: ( x in b . i implies ( x <> (S . 0) `1 & [x,((S . 0) `1)] in the InternalRel of R ) ) assume A60: x in b . i ; ::_thesis: ( x <> (S . 0) `1 & [x,((S . 0) `1)] in the InternalRel of R ) A61: a0 = [a0,b0] `1 ; b0 = [a0,b0] `2 .= b by A53, A54 ; then A62: x in (b0t . i) \ {a0} by A51, A60; then A63: x in b0t . i ; not x in {a0} by A62, XBOOLE_0:def_5; hence A64: x <> (S . 0) `1 by A61, A53, A54, TARSKI:def_1; ::_thesis: [x,((S . 0) `1)] in the InternalRel of R b . i c= the carrier of R by FINSUB_1:def_5; then reconsider x9 = x as Element of R by A60; A65: x in S02 . (i + i0) by A50, A63, NAT_1:def_3; consider S02ib being Element of Fin the carrier of R such that A66: S02ib = S02 . (i + i0) and A67: S02ib <> {} and A68: S0max . (i + i0) = PosetMax S02ib by A49; PosetMax S02ib = a0 by A52, A68, NAT_1:11; then a0 is_maximal_wrt S02ib, the InternalRel of R by A67, Def13; then not [a0,x] in the InternalRel of R by A61, A64, A65, A66, A53, A54, WAYBEL_4:def_23; then not a0 <= x9 by ORDERS_2:def_5; then x9 <= a0 by WAYBEL_0:def_29; hence [x,((S . 0) `1)] in the InternalRel of R by A61, A53, A54, ORDERS_2:def_5; ::_thesis: verum end; A69: for n being Nat st S2[n] holds S2[n + 1] proof let n be Nat; ::_thesis: ( S2[n] implies S2[n + 1] ) assume S2[n] ; ::_thesis: S2[n + 1] let i be Nat; ::_thesis: ex b being sequence of (FinPoset R) st ( b = (S . (n + 1)) `2 & ( for x being set st x in b . i holds ( x <> (S . (n + 1)) `1 & [x,((S . (n + 1)) `1)] in the InternalRel of R ) ) ) set n1 = n + 1; reconsider n1 = n + 1 as Nat ; set b = (S . n1) `2 ; consider Sn2 being sequence of (FinPoset R), Smax being Function of NAT, the carrier of R, an being Element of R, ix being Nat, bnt, bn being sequence of (FinPoset R) such that Sn2 = (S . n) `2 and A70: for i being Nat ex Sn2i being Element of Fin the carrier of R st ( Sn2i = Sn2 . i & Sn2i <> {} & Smax . i = PosetMax Sn2i ) and an = MinElement ((rng Smax),R) and ix = Smax mindex an and A71: bnt = Sn2 ^\ ix and A72: for i being Nat holds bn . i = (bnt . i) \ {an} and A73: for i being Nat st ix <= i holds Smax . i = an and A74: S . (n + 1) = [an,bn] by A56; (S . n1) `2 in zz ; then ex z being sequence of (FinPoset R) st ( z = (S . n1) `2 & z is descending ) ; then reconsider b = (S . n1) `2 as sequence of (FinPoset R) ; take b ; ::_thesis: ( b = (S . (n + 1)) `2 & ( for x being set st x in b . i holds ( x <> (S . (n + 1)) `1 & [x,((S . (n + 1)) `1)] in the InternalRel of R ) ) ) thus b = (S . (n + 1)) `2 ; ::_thesis: for x being set st x in b . i holds ( x <> (S . (n + 1)) `1 & [x,((S . (n + 1)) `1)] in the InternalRel of R ) let x be set ; ::_thesis: ( x in b . i implies ( x <> (S . (n + 1)) `1 & [x,((S . (n + 1)) `1)] in the InternalRel of R ) ) assume A75: x in b . i ; ::_thesis: ( x <> (S . (n + 1)) `1 & [x,((S . (n + 1)) `1)] in the InternalRel of R ) A76: an = [an,bn] `1 ; bn = [an,bn] `2 .= b by A74 ; then A77: x in (bnt . i) \ {an} by A72, A75; then A78: x in bnt . i ; not x in {an} by A77, XBOOLE_0:def_5; hence A79: x <> (S . (n + 1)) `1 by A76, A74, TARSKI:def_1; ::_thesis: [x,((S . (n + 1)) `1)] in the InternalRel of R b . i c= the carrier of R by FINSUB_1:def_5; then reconsider x9 = x as Element of R by A75; A80: x in Sn2 . (i + ix) by A71, A78, NAT_1:def_3; consider Sn2ib being Element of Fin the carrier of R such that A81: Sn2ib = Sn2 . (i + ix) and A82: Sn2ib <> {} and A83: Smax . (i + ix) = PosetMax Sn2ib by A70; PosetMax Sn2ib = an by A73, A83, NAT_1:11; then an is_maximal_wrt Sn2ib, the InternalRel of R by A82, Def13; then not [an,x] in the InternalRel of R by A76, A79, A80, A81, A74, WAYBEL_4:def_23; then not an <= x9 by ORDERS_2:def_5; then x9 <= an by WAYBEL_0:def_29; hence [x,((S . (n + 1)) `1)] in the InternalRel of R by A76, A74, ORDERS_2:def_5; ::_thesis: verum end; A84: for n being Nat holds S2[n] from NAT_1:sch_2(A59, A69); defpred S3[ Nat] means ex b being sequence of (FinPoset R) ex i being Nat st ( b = (S . $1) `2 & a . ($1 + 1) in b . i ); A85: S3[ 0 ] proof set b = (S . 0) `2 ; (S . 0) `2 in zz ; then ex z being sequence of (FinPoset R) st ( z = (S . 0) `2 & z is descending ) ; then reconsider b = (S . 0) `2 as sequence of (FinPoset R) ; take b ; ::_thesis: ex i being Nat st ( b = (S . 0) `2 & a . (0 + 1) in b . i ) A86: a . (0 + 1) = (S . (0 + 1)) `1 by A58; consider S12 being sequence of (FinPoset R), S1max being Function of NAT, the carrier of R, a1 being Element of R, i1 being Nat, b1t, b1 being sequence of (FinPoset R) such that A87: S12 = (S . 0) `2 and A88: for i being Nat ex S12i being Element of Fin the carrier of R st ( S12i = S12 . i & S12i <> {} & S1max . i = PosetMax S12i ) and A89: a1 = MinElement ((rng S1max),R) and i1 = S1max mindex a1 and b1t = S12 ^\ i1 and for i being Nat holds b1 . i = (b1t . i) \ {a1} and for i being Nat st i1 <= i holds S1max . i = a1 and A90: S . (0 + 1) = [a1,b1] by A55; rng S1max c= the carrier of R by RELAT_1:def_19; then a1 in rng S1max by A1, A89, Def17; then consider i being set such that A91: i in dom S1max and A92: S1max . i = a1 by FUNCT_1:def_3; A93: ex S12i being Element of Fin the carrier of R st ( S12i = S12 . i & S12i <> {} & S1max . i = PosetMax S12i ) by A88, A91; reconsider i = i as Nat by A91; take i ; ::_thesis: ( b = (S . 0) `2 & a . (0 + 1) in b . i ) thus b = (S . 0) `2 ; ::_thesis: a . (0 + 1) in b . i A94: a1 in b . i by A87, A92, A93, Def13; [a1,b1] `1 = a1 ; hence a . (0 + 1) in b . i by A86, A90, A94; ::_thesis: verum end; A95: for n being Nat st S3[n] holds S3[n + 1] proof let n be Nat; ::_thesis: ( S3[n] implies S3[n + 1] ) assume S3[n] ; ::_thesis: S3[n + 1] set b = (S . (n + 1)) `2 ; (S . (n + 1)) `2 in zz ; then ex z being sequence of (FinPoset R) st ( z = (S . (n + 1)) `2 & z is descending ) ; then reconsider b = (S . (n + 1)) `2 as sequence of (FinPoset R) ; take b ; ::_thesis: ex i being Nat st ( b = (S . (n + 1)) `2 & a . ((n + 1) + 1) in b . i ) set n1 = n + 1; reconsider n1 = n + 1 as Nat ; consider Sn12 being sequence of (FinPoset R), Sn1max being Function of NAT, the carrier of R, an1 being Element of R, in1 being Nat, bn1t, bn1 being sequence of (FinPoset R) such that A96: Sn12 = (S . n1) `2 and A97: for i being Nat ex Sn12i being Element of Fin the carrier of R st ( Sn12i = Sn12 . i & Sn12i <> {} & Sn1max . i = PosetMax Sn12i ) and A98: an1 = MinElement ((rng Sn1max),R) and in1 = Sn1max mindex an1 and bn1t = Sn12 ^\ in1 and for i being Nat holds bn1 . i = (bn1t . i) \ {an1} and for i being Nat st in1 <= i holds Sn1max . i = an1 and A99: S . (n1 + 1) = [an1,bn1] by A55; rng Sn1max c= the carrier of R by RELAT_1:def_19; then an1 in rng Sn1max by A1, A98, Def17; then consider i being set such that A100: i in dom Sn1max and A101: Sn1max . i = an1 by FUNCT_1:def_3; A102: ex Sn12i being Element of Fin the carrier of R st ( Sn12i = Sn12 . i & Sn12i <> {} & Sn1max . i = PosetMax Sn12i ) by A97, A100; reconsider i = i as Nat by A100; take i ; ::_thesis: ( b = (S . (n + 1)) `2 & a . ((n + 1) + 1) in b . i ) thus b = (S . (n + 1)) `2 ; ::_thesis: a . ((n + 1) + 1) in b . i A103: an1 in b . i by A96, A101, A102, Def13; [an1,bn1] `1 = an1 ; hence a . ((n + 1) + 1) in b . i by A58, A103, A99; ::_thesis: verum end; A104: for n being Nat holds S3[n] from NAT_1:sch_2(A85, A95); defpred S4[ Nat] means ( a . ($1 + 1) <> a . $1 & [(a . ($1 + 1)),(a . $1)] in the InternalRel of R ); A105: S4[ 0 ] proof A106: a . 0 = (S . 0) `1 by A58; consider b being sequence of (FinPoset R), i being Nat such that A107: b = (S . 0) `2 and A108: a . (0 + 1) in b . i by A85; ex bb being sequence of (FinPoset R) st ( bb = (S . 0) `2 & ( for x being set st x in bb . i holds ( x <> (S . 0) `1 & [x,((S . 0) `1)] in the InternalRel of R ) ) ) by A59; hence ( a . (0 + 1) <> a . 0 & [(a . (0 + 1)),(a . 0)] in the InternalRel of R ) by A106, A107, A108; ::_thesis: verum end; A109: for n being Nat st S4[n] holds S4[n + 1] proof let n be Nat; ::_thesis: ( S4[n] implies S4[n + 1] ) assume that a . (n + 1) <> a . n and [(a . (n + 1)),(a . n)] in the InternalRel of R ; ::_thesis: S4[n + 1] A110: a . (n + 1) = (S . (n + 1)) `1 by A58; consider b being sequence of (FinPoset R), i being Nat such that A111: b = (S . (n + 1)) `2 and A112: a . ((n + 1) + 1) in b . i by A104; ex bb being sequence of (FinPoset R) st ( bb = (S . (n + 1)) `2 & ( for x being set st x in bb . i holds ( x <> (S . (n + 1)) `1 & [x,((S . (n + 1)) `1)] in the InternalRel of R ) ) ) by A84; hence S4[n + 1] by A110, A111, A112; ::_thesis: verum end; for n being Nat holds S4[n] from NAT_1:sch_2(A105, A109); then a is descending by WELLFND1:def_6; hence contradiction by A1, WELLFND1:14; ::_thesis: verum end;