:: SCPISORT semantic presentation begin definition let f be FinSequence of INT ; let s be State of SCMPDS; let m be Element of NAT ; predf is_FinSequence_on s,m means :Def1: :: SCPISORT:def 1 for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)); end; :: deftheorem Def1 defines is_FinSequence_on SCPISORT:def_1_:_ for f being FinSequence of INT for s being State of SCMPDS for m being Element of NAT holds ( f is_FinSequence_on s,m iff for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) ); theorem Th1: :: SCPISORT:1 for s being State of SCMPDS for n, m being Element of NAT ex f being FinSequence of INT st ( len f = n & ( for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) ) ) proof let s be State of SCMPDS; ::_thesis: for n, m being Element of NAT ex f being FinSequence of INT st ( len f = n & ( for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) ) ) let n, m be Element of NAT ; ::_thesis: ex f being FinSequence of INT st ( len f = n & ( for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) ) ) deffunc H1( Nat) -> set = s . (intpos (m + $1)); consider f being FinSequence such that A1: ( len f = n & ( for i being Nat st i in dom f holds f . i = H1(i) ) ) from FINSEQ_1:sch_2(); now__::_thesis:_for_i_being_Nat_st_i_in_dom_f_holds_ f_._i_in_INT let i be Nat; ::_thesis: ( i in dom f implies f . i in INT ) reconsider a = i as Element of NAT by ORDINAL1:def_12; assume i in dom f ; ::_thesis: f . i in INT then f . i = s . (intpos (m + a)) by A1; hence f . i in INT by INT_1:def_2; ::_thesis: verum end; then reconsider f = f as FinSequence of INT by FINSEQ_2:12; take f ; ::_thesis: ( len f = n & ( for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) ) ) thus len f = n by A1; ::_thesis: for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) hereby ::_thesis: verum let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len f implies f . i = s . (intpos (m + i)) ) assume ( 1 <= i & i <= len f ) ; ::_thesis: f . i = s . (intpos (m + i)) then i in dom f by FINSEQ_3:25; hence f . i = s . (intpos (m + i)) by A1; ::_thesis: verum end; end; theorem :: SCPISORT:2 for s being State of SCMPDS for n, m being Element of NAT ex f being FinSequence of INT st ( len f = n & f is_FinSequence_on s,m ) proof let s be State of SCMPDS; ::_thesis: for n, m being Element of NAT ex f being FinSequence of INT st ( len f = n & f is_FinSequence_on s,m ) let n, m be Element of NAT ; ::_thesis: ex f being FinSequence of INT st ( len f = n & f is_FinSequence_on s,m ) consider f being FinSequence of INT such that A1: len f = n and A2: for i being Element of NAT st 1 <= i & i <= len f holds f . i = s . (intpos (m + i)) by Th1; take f ; ::_thesis: ( len f = n & f is_FinSequence_on s,m ) thus len f = n by A1; ::_thesis: f is_FinSequence_on s,m thus f is_FinSequence_on s,m by A2, Def1; ::_thesis: verum end; theorem Th3: :: SCPISORT:3 for f, g being FinSequence of INT for m, n being Element of NAT st 1 <= n & n <= len f & 1 <= m & m <= len f & len f = len g & f . m = g . n & f . n = g . m & ( for k being Element of NAT st k <> m & k <> n & 1 <= k & k <= len f holds f . k = g . k ) holds f,g are_fiberwise_equipotent proof let f, g be FinSequence of INT ; ::_thesis: for m, n being Element of NAT st 1 <= n & n <= len f & 1 <= m & m <= len f & len f = len g & f . m = g . n & f . n = g . m & ( for k being Element of NAT st k <> m & k <> n & 1 <= k & k <= len f holds f . k = g . k ) holds f,g are_fiberwise_equipotent let m, n be Element of NAT ; ::_thesis: ( 1 <= n & n <= len f & 1 <= m & m <= len f & len f = len g & f . m = g . n & f . n = g . m & ( for k being Element of NAT st k <> m & k <> n & 1 <= k & k <= len f holds f . k = g . k ) implies f,g are_fiberwise_equipotent ) assume that A1: ( 1 <= n & n <= len f ) and A2: ( 1 <= m & m <= len f ) and A3: len f = len g and A4: ( f . m = g . n & f . n = g . m ) and A5: for k being Element of NAT st k <> m & k <> n & 1 <= k & k <= len f holds f . k = g . k ; ::_thesis: f,g are_fiberwise_equipotent A6: m in Seg (len f) by A2, FINSEQ_1:1; A7: Seg (len f) = dom f by FINSEQ_1:def_3; A8: now__::_thesis:_for_k_being_set_st_k_<>_m_&_k_<>_n_&_k_in_dom_f_holds_ f_._k_=_g_._k let k be set ; ::_thesis: ( k <> m & k <> n & k in dom f implies f . k = g . k ) assume that A9: ( k <> m & k <> n ) and A10: k in dom f ; ::_thesis: f . k = g . k reconsider i = k as Element of NAT by A10; ( 1 <= i & i <= len f ) by A7, A10, FINSEQ_1:1; hence f . k = g . k by A5, A9; ::_thesis: verum end; ( n in dom f & dom f = dom g ) by A1, A3, A7, FINSEQ_1:1, FINSEQ_1:def_3; hence f,g are_fiberwise_equipotent by A4, A7, A6, A8, RFINSEQ:28; ::_thesis: verum end; set A = NAT ; set D = SCM-Data-Loc ; theorem Th4: :: SCPISORT:4 for s1, s2 being State of SCMPDS st ( for a being Int_position holds s1 . a = s2 . a ) holds Initialize s1 = Initialize s2 proof let s1, s2 be State of SCMPDS; ::_thesis: ( ( for a being Int_position holds s1 . a = s2 . a ) implies Initialize s1 = Initialize s2 ) assume for a being Int_position holds s1 . a = s2 . a ; ::_thesis: Initialize s1 = Initialize s2 then DataPart s1 = DataPart s2 by SCMPDS_4:8; hence Initialize s1 = Initialize s2 by MEMSTR_0:80; ::_thesis: verum end; theorem Th5: :: SCPISORT:5 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being halt-free Program of SCMPDS for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being halt-free Program of SCMPDS for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) let s be State of SCMPDS; ::_thesis: for I being halt-free Program of SCMPDS for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) let I be halt-free Program of SCMPDS; ::_thesis: for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) let j be shiftable parahalting Instruction of SCMPDS; ::_thesis: ( I is_closed_on s,P & I is_halting_on s,P implies ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) ) set Mj = Load j; A1: ( Load j is_closed_on IExec (I,P,(Initialize s)),P & Load j is_halting_on IExec (I,P,(Initialize s)),P ) by SCMPDS_6:20, SCMPDS_6:21; assume ( I is_closed_on s,P & I is_halting_on s,P ) ; ::_thesis: ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) then ( I ';' (Load j) is_closed_on s,P & I ';' (Load j) is_halting_on s,P ) by A1, SCMPDS_7:24; hence ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) by SCMPDS_4:def_3; ::_thesis: verum end; theorem :: SCPISORT:6 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free Program of SCMPDS for J being parahalting shiftable Program of SCMPDS for a being Int_position st I is_closed_on s,P & I is_halting_on s,P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for I being halt-free Program of SCMPDS for J being parahalting shiftable Program of SCMPDS for a being Int_position st I is_closed_on s,P & I is_halting_on s,P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free Program of SCMPDS for J being parahalting shiftable Program of SCMPDS for a being Int_position st I is_closed_on s,P & I is_halting_on s,P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let I be halt-free Program of SCMPDS; ::_thesis: for J being parahalting shiftable Program of SCMPDS for a being Int_position st I is_closed_on s,P & I is_halting_on s,P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let J be parahalting shiftable Program of SCMPDS; ::_thesis: for a being Int_position st I is_closed_on s,P & I is_halting_on s,P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let a be Int_position; ::_thesis: ( I is_closed_on s,P & I is_halting_on s,P implies (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a ) ( J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P ) by SCMPDS_6:20, SCMPDS_6:21; hence ( I is_closed_on s,P & I is_halting_on s,P implies (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a ) by SCMPDS_7:30; ::_thesis: verum end; theorem :: SCPISORT:7 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free parahalting Program of SCMPDS for J being shiftable Program of SCMPDS for a being Int_position st J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for I being halt-free parahalting Program of SCMPDS for J being shiftable Program of SCMPDS for a being Int_position st J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free parahalting Program of SCMPDS for J being shiftable Program of SCMPDS for a being Int_position st J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let I be halt-free parahalting Program of SCMPDS; ::_thesis: for J being shiftable Program of SCMPDS for a being Int_position st J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let J be shiftable Program of SCMPDS; ::_thesis: for a being Int_position st J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a let a be Int_position; ::_thesis: ( J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P implies (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a ) ( I is_closed_on s,P & I is_halting_on s,P ) by SCMPDS_6:20, SCMPDS_6:21; hence ( J is_closed_on IExec (I,P,s),P & J is_halting_on IExec (I,P,s),P implies (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(Initialize (IExec (I,P,s))))) . a ) by SCMPDS_7:30; ::_thesis: verum end; theorem :: SCPISORT:8 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being Program of SCMPDS for J being parahalting shiftable Program of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being Program of SCMPDS for J being parahalting shiftable Program of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) let s be State of SCMPDS; ::_thesis: for I being Program of SCMPDS for J being parahalting shiftable Program of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) let I be Program of SCMPDS; ::_thesis: for J being parahalting shiftable Program of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) let J be parahalting shiftable Program of SCMPDS; ::_thesis: ( I is_closed_on s,P & I is_halting_on s,P implies ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) ) A1: ( J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P ) by SCMPDS_6:20, SCMPDS_6:21; assume ( I is_closed_on s,P & I is_halting_on s,P ) ; ::_thesis: ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) hence ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) by A1, SCMPDS_7:24; ::_thesis: verum end; theorem :: SCPISORT:9 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being parahalting Program of SCMPDS for J being shiftable Program of SCMPDS st J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being parahalting Program of SCMPDS for J being shiftable Program of SCMPDS st J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) let s be State of SCMPDS; ::_thesis: for I being parahalting Program of SCMPDS for J being shiftable Program of SCMPDS st J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) let I be parahalting Program of SCMPDS; ::_thesis: for J being shiftable Program of SCMPDS st J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P holds ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) let J be shiftable Program of SCMPDS; ::_thesis: ( J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P implies ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) ) A1: ( I is_closed_on s,P & I is_halting_on s,P ) by SCMPDS_6:20, SCMPDS_6:21; assume ( J is_closed_on IExec (I,P,(Initialize s)),P & J is_halting_on IExec (I,P,(Initialize s)),P ) ; ::_thesis: ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) hence ( I ';' J is_closed_on s,P & I ';' J is_halting_on s,P ) by A1, SCMPDS_7:24; ::_thesis: verum end; theorem :: SCPISORT:10 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being Program of SCMPDS for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being Program of SCMPDS for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) let s be State of SCMPDS; ::_thesis: for I being Program of SCMPDS for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) let I be Program of SCMPDS; ::_thesis: for j being shiftable parahalting Instruction of SCMPDS st I is_closed_on s,P & I is_halting_on s,P holds ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) let j be shiftable parahalting Instruction of SCMPDS; ::_thesis: ( I is_closed_on s,P & I is_halting_on s,P implies ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) ) A1: ( Load j is_closed_on IExec (I,P,(Initialize s)),P & Load j is_halting_on IExec (I,P,(Initialize s)),P ) by SCMPDS_6:20, SCMPDS_6:21; assume ( I is_closed_on s,P & I is_halting_on s,P ) ; ::_thesis: ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) then ( I ';' (Load j) is_closed_on s,P & I ';' (Load j) is_halting_on s,P ) by A1, SCMPDS_7:24; hence ( I ';' j is_closed_on s,P & I ';' j is_halting_on s,P ) by SCMPDS_4:def_3; ::_thesis: verum end; Lm1: for a being Int_position for i being Integer for n being Element of NAT for I being Program of SCMPDS holds card (stop (for-down (a,i,n,I))) = (card I) + 4 proof let a be Int_position; ::_thesis: for i being Integer for n being Element of NAT for I being Program of SCMPDS holds card (stop (for-down (a,i,n,I))) = (card I) + 4 let i be Integer; ::_thesis: for n being Element of NAT for I being Program of SCMPDS holds card (stop (for-down (a,i,n,I))) = (card I) + 4 let n be Element of NAT ; ::_thesis: for I being Program of SCMPDS holds card (stop (for-down (a,i,n,I))) = (card I) + 4 let I be Program of SCMPDS; ::_thesis: card (stop (for-down (a,i,n,I))) = (card I) + 4 thus card (stop (for-down (a,i,n,I))) = (card (for-down (a,i,n,I))) + 1 by COMPOS_1:55 .= ((card I) + 3) + 1 by SCMPDS_7:41 .= (card I) + 4 ; ::_thesis: verum end; Lm2: for a being Int_position for i being Integer for n being Element of NAT for I being Program of SCMPDS holds for-down (a,i,n,I) = ((a,i) <=0_goto ((card I) + 3)) ';' ((I ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2)))) proof let a be Int_position; ::_thesis: for i being Integer for n being Element of NAT for I being Program of SCMPDS holds for-down (a,i,n,I) = ((a,i) <=0_goto ((card I) + 3)) ';' ((I ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2)))) let i be Integer; ::_thesis: for n being Element of NAT for I being Program of SCMPDS holds for-down (a,i,n,I) = ((a,i) <=0_goto ((card I) + 3)) ';' ((I ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2)))) let n be Element of NAT ; ::_thesis: for I being Program of SCMPDS holds for-down (a,i,n,I) = ((a,i) <=0_goto ((card I) + 3)) ';' ((I ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2)))) let I be Program of SCMPDS; ::_thesis: for-down (a,i,n,I) = ((a,i) <=0_goto ((card I) + 3)) ';' ((I ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2)))) set i1 = (a,i) <=0_goto ((card I) + 3); set i2 = AddTo (a,i,(- n)); set i3 = goto (- ((card I) + 2)); thus for-down (a,i,n,I) = ((((a,i) <=0_goto ((card I) + 3)) ';' I) ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2))) by SCMPDS_7:def_2 .= ((a,i) <=0_goto ((card I) + 3)) ';' ((I ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2)))) by SCMPDS_7:2 ; ::_thesis: verum end; Lm3: for I being Program of SCMPDS for a being Int_position for i being Integer for n being Element of NAT holds Shift ((I ';' (AddTo (a,i,(- n)))),1) c= for-down (a,i,n,I) proof let I be Program of SCMPDS; ::_thesis: for a being Int_position for i being Integer for n being Element of NAT holds Shift ((I ';' (AddTo (a,i,(- n)))),1) c= for-down (a,i,n,I) let a be Int_position; ::_thesis: for i being Integer for n being Element of NAT holds Shift ((I ';' (AddTo (a,i,(- n)))),1) c= for-down (a,i,n,I) let i be Integer; ::_thesis: for n being Element of NAT holds Shift ((I ';' (AddTo (a,i,(- n)))),1) c= for-down (a,i,n,I) let n be Element of NAT ; ::_thesis: Shift ((I ';' (AddTo (a,i,(- n)))),1) c= for-down (a,i,n,I) set i1 = (a,i) <=0_goto ((card I) + 3); set i2 = AddTo (a,i,(- n)); set i3 = goto (- ((card I) + 2)); A1: for-down (a,i,n,I) = ((((a,i) <=0_goto ((card I) + 3)) ';' I) ';' (AddTo (a,i,(- n)))) ';' (goto (- ((card I) + 2))) by SCMPDS_7:def_2 .= (((a,i) <=0_goto ((card I) + 3)) ';' (I ';' (AddTo (a,i,(- n))))) ';' (goto (- ((card I) + 2))) by SCMPDS_4:15 .= ((Load ((a,i) <=0_goto ((card I) + 3))) ';' (I ';' (AddTo (a,i,(- n))))) ';' (goto (- ((card I) + 2))) by SCMPDS_4:def_2 .= ((Load ((a,i) <=0_goto ((card I) + 3))) ';' (I ';' (AddTo (a,i,(- n))))) ';' (Load (goto (- ((card I) + 2)))) by SCMPDS_4:def_3 ; card (Load ((a,i) <=0_goto ((card I) + 3))) = 1 by COMPOS_1:54; hence Shift ((I ';' (AddTo (a,i,(- n)))),1) c= for-down (a,i,n,I) by A1, SCMPDS_7:3; ::_thesis: verum end; begin scheme :: SCPISORT:sch 1 ForDownHalt{ P1[ set ], F1() -> 0 -started State of SCMPDS, F2() -> Instruction-Sequence of SCMPDS, F3() -> halt-free shiftable Program of SCMPDS, F4() -> Int_position, F5() -> Integer, F6() -> Element of NAT } : ( for-down (F4(),F5(),F6(),F3()) is_closed_on F1(),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on F1(),F2() ) provided A1: F6() > 0 and A2: P1[F1()] and A3: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F4() = F1() . F4() & t . (DataLoc ((F1() . F4()),F5())) > 0 holds ( (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() & (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc ((F1() . F4()),F5())) = (t . (DataLoc ((F1() . F4()),F5()))) - F6() & F3() is_closed_on t,Q & F3() is_halting_on t,Q & P1[ Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t))] ) proof set i1 = (F4(),F5()) <=0_goto ((card F3()) + 3); set J = F3() ';' (AddTo (F4(),F5(),(- F6()))); set i3 = goto (- ((card F3()) + 2)); set b = DataLoc ((F1() . F4()),F5()); set FOR = for-down (F4(),F5(),F6(),F3()); set pFOR = stop (for-down (F4(),F5(),F6(),F3())); set pJ = stop (F3() ';' (AddTo (F4(),F5(),(- F6())))); defpred S1[ Nat] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . (DataLoc ((F1() . F4()),F5())) <= $1 & P1[t] & t . F4() = F1() . F4() holds ( for-down (F4(),F5(),F6(),F3()) is_closed_on t,Q & for-down (F4(),F5(),F6(),F3()) is_halting_on t,Q ); A4: for k being Element of NAT st S1[k] holds S1[k + 1] proof let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A5: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_t_._(DataLoc_((F1()_._F4()),F5()))_<=_k_+_1_&_P1[t]_&_t_._F4()_=_F1()_._F4()_holds_ (_for-down_(F4(),F5(),F6(),F3())_is_closed_on_t,Q_&_for-down_(F4(),F5(),F6(),F3())_is_halting_on_t,Q_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (DataLoc ((F1() . F4()),F5())) <= k + 1 & P1[t] & t . F4() = F1() . F4() holds ( for-down (F4(),F5(),F6(),F3()) is_closed_on b2,b3 & for-down (F4(),F5(),F6(),F3()) is_halting_on b2,b3 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (DataLoc ((F1() . F4()),F5())) <= k + 1 & P1[t] & t . F4() = F1() . F4() implies ( for-down (F4(),F5(),F6(),F3()) is_closed_on b1,b2 & for-down (F4(),F5(),F6(),F3()) is_halting_on b1,b2 ) ) A6: Initialize t = t by MEMSTR_0:44; assume A7: t . (DataLoc ((F1() . F4()),F5())) <= k + 1 ; ::_thesis: ( P1[t] & t . F4() = F1() . F4() implies ( for-down (F4(),F5(),F6(),F3()) is_closed_on b1,b2 & for-down (F4(),F5(),F6(),F3()) is_halting_on b1,b2 ) ) assume A8: P1[t] ; ::_thesis: ( t . F4() = F1() . F4() implies ( for-down (F4(),F5(),F6(),F3()) is_closed_on b1,b2 & for-down (F4(),F5(),F6(),F3()) is_halting_on b1,b2 ) ) assume A9: t . F4() = F1() . F4() ; ::_thesis: ( for-down (F4(),F5(),F6(),F3()) is_closed_on b1,b2 & for-down (F4(),F5(),F6(),F3()) is_halting_on b1,b2 ) percases ( t . (DataLoc ((F1() . F4()),F5())) <= 0 or t . (DataLoc ((F1() . F4()),F5())) > 0 ) ; suppose t . (DataLoc ((F1() . F4()),F5())) <= 0 ; ::_thesis: ( for-down (F4(),F5(),F6(),F3()) is_closed_on b1,b2 & for-down (F4(),F5(),F6(),F3()) is_halting_on b1,b2 ) hence ( for-down (F4(),F5(),F6(),F3()) is_closed_on t,Q & for-down (F4(),F5(),F6(),F3()) is_halting_on t,Q ) by A9, SCMPDS_7:44; ::_thesis: verum end; supposeA10: t . (DataLoc ((F1() . F4()),F5())) > 0 ; ::_thesis: ( for-down (F4(),F5(),F6(),F3()) is_closed_on b1,b2 & for-down (F4(),F5(),F6(),F3()) is_halting_on b1,b2 ) A11: 0 in dom (stop (for-down (F4(),F5(),F6(),F3()))) by COMPOS_1:36; - (- F6()) > 0 by A1; then - F6() < 0 ; then - F6() <= - 1 by INT_1:8; then A12: (- F6()) + (t . (DataLoc ((F1() . F4()),F5()))) <= (- 1) + (t . (DataLoc ((F1() . F4()),F5()))) by XREAL_1:6; (t . (DataLoc ((F1() . F4()),F5()))) - 1 <= k by A7, XREAL_1:20; then A13: (- F6()) + (t . (DataLoc ((F1() . F4()),F5()))) <= k by A12, XXREAL_0:2; set Q2 = Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))); set Q3 = Q +* (stop (for-down (F4(),F5(),F6(),F3()))); set t4 = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1); set Q4 = Q +* (stop (for-down (F4(),F5(),F6(),F3()))); set Jt = IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t); A14: stop (F3() ';' (AddTo (F4(),F5(),(- F6())))) c= Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by FUNCT_4:25; A15: for-down (F4(),F5(),F6(),F3()) = ((F4(),F5()) <=0_goto ((card F3()) + 3)) ';' ((F3() ';' (AddTo (F4(),F5(),(- F6())))) ';' (goto (- ((card F3()) + 2)))) by Lm2; A16: Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(0 + 1)) = Following ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,0))) by EXTPRO_1:3 .= Following ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t) by EXTPRO_1:2 .= Exec (((F4(),F5()) <=0_goto ((card F3()) + 3)),t) by A15, A6, SCMPDS_6:11 ; for x being Int_position holds t . x = (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)) . x by A16, SCMPDS_2:56; then A17: DataPart t = DataPart (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)) by SCMPDS_4:8; A18: (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc ((F1() . F4()),F5())) = (t . (DataLoc ((F1() . F4()),F5()))) - F6() by A3, A8, A9, A10; A19: (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() by A3, A8, A9, A10; set m2 = LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t); set t5 = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t))); set Q5 = Q +* (stop (for-down (F4(),F5(),F6(),F3()))); set l1 = (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1; A20: card (F3() ';' (AddTo (F4(),F5(),(- F6())))) = (card F3()) + 1 by SCMP_GCD:4; A21: t . (DataLoc ((t . F4()),F5())) = t . (DataLoc ((F1() . F4()),F5())) by A9; A22: IC t = 0 by A6, MEMSTR_0:47; A23: ( F3() is_closed_on t,Q & F3() is_halting_on t,Q ) by A3, A8, A9, A10; then A24: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_closed_on t,Q by Th5; then A25: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_closed_on t,Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by A6, SCMPDS_6:24; (card F3()) + 2 < (card F3()) + 3 by XREAL_1:6; then A26: (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 in dom (for-down (F4(),F5(),F6(),F3())) by A20, SCMPDS_7:42; set m3 = (LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1; set t6 = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1)); set Q6 = Q +* (stop (for-down (F4(),F5(),F6(),F3()))); A27: Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1)) = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t))) by EXTPRO_1:4; A28: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_halting_on t,Q by A23, Th5; then A29: Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) halts_on t by A6, SCMPDS_6:def_3; (Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))) +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) halts_on t by A28, A6, SCMPDS_6:def_3; then A30: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_halting_on t,Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by A6, SCMPDS_6:def_3; set m4 = ((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1; set t7 = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)); set Q7 = Q +* (stop (for-down (F4(),F5(),F6(),F3()))); A31: stop (for-down (F4(),F5(),F6(),F3())) c= Q +* (stop (for-down (F4(),F5(),F6(),F3()))) by FUNCT_4:25; A32: for-down (F4(),F5(),F6(),F3()) c= stop (for-down (F4(),F5(),F6(),F3())) by AFINSQ_1:74; then A33: for-down (F4(),F5(),F6(),F3()) c= Q +* (stop (for-down (F4(),F5(),F6(),F3()))) by A31, XBOOLE_1:1; Shift ((F3() ';' (AddTo (F4(),F5(),(- F6())))),1) c= for-down (F4(),F5(),F6(),F3()) by Lm3; then Shift ((F3() ';' (AddTo (F4(),F5(),(- F6())))),1) c= stop (for-down (F4(),F5(),F6(),F3())) by A32, XBOOLE_1:1; then A34: Shift ((F3() ';' (AddTo (F4(),F5(),(- F6())))),1) c= Q +* (stop (for-down (F4(),F5(),F6(),F3()))) by A31, XBOOLE_1:1; A35: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)) = succ (IC t) by A10, A16, A21, SCMPDS_2:56 .= 0 + 1 by A22 ; then A36: DataPart (Comput ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t,(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) = DataPart (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) by A14, A30, A25, A17, A34, SCMPDS_7:18; then A37: DataPart (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) = DataPart (Result ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) by A29, EXTPRO_1:23 .= DataPart (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) by SCMPDS_4:def_5 ; A38: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) = (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 by A14, A30, A25, A35, A17, A34, SCMPDS_7:18; then A39: CurInstr ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1)))) = (Q +* (stop (for-down (F4(),F5(),F6(),F3())))) . ((card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1) by A27, PBOOLE:143 .= (for-down (F4(),F5(),F6(),F3())) . ((card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1) by A26, A33, GRFUNC_1:2 .= goto (- ((card F3()) + 2)) by A20, SCMPDS_7:43 ; A40: Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)) = Following ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card F3()) + 2))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1)))) by A39 ; IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1))),(0 - ((card F3()) + 2))) by A40, SCMPDS_2:54 .= 0 by A20, A38, A27, SCMPDS_7:1 ; then A41: Initialize (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))) = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)) by MEMSTR_0:46; InsCode (goto (- ((card F3()) + 2))) = 14 by SCMPDS_2:12; then InsCode (goto (- ((card F3()) + 2))) in {0,4,5,6,14} by ENUMSET1:def_3; then Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)) = Initialize (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1))) by A40, A41, SCMPDS_8:3 .= Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) by A37, A27, MEMSTR_0:80 ; then A42: P1[ Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))] by A3, A8, A9, A10; A43: (Q +* (stop (for-down (F4(),F5(),F6(),F3())))) +* (stop (for-down (F4(),F5(),F6(),F3()))) = Q +* (stop (for-down (F4(),F5(),F6(),F3()))) ; (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) . (DataLoc ((F1() . F4()),F5())) = (Comput ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t,(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) . (DataLoc ((F1() . F4()),F5())) by A36, SCMPDS_4:8 .= (Result ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) . (DataLoc ((F1() . F4()),F5())) by A29, EXTPRO_1:23 .= (t . (DataLoc ((F1() . F4()),F5()))) - F6() by A18, SCMPDS_4:def_5 ; then A44: (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))) . (DataLoc ((F1() . F4()),F5())) = (- F6()) + (t . (DataLoc ((F1() . F4()),F5()))) by A27, A40, SCMPDS_2:54; (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) . F4() = (Comput ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t,(LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)))) . F4() by A36, SCMPDS_4:8 .= (Result ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) . F4() by A29, EXTPRO_1:23 .= F1() . F4() by A9, A19, SCMPDS_4:def_5 ; then A45: (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))) . F4() = F1() . F4() by A27, A40, SCMPDS_2:54; then A46: for-down (F4(),F5(),F6(),F3()) is_closed_on Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)),Q +* (stop (for-down (F4(),F5(),F6(),F3()))) by A5, A42, A44, A13, A41; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(for-down_(F4(),F5(),F6(),F3())))),t,k))_in_dom_(stop_(for-down_(F4(),F5(),F6(),F3()))) let k be Element of NAT ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,b1)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) percases ( k < ((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,b1)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) then A47: k <= (LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1 by INT_1:7; hereby ::_thesis: verum percases ( k <= LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t) or k = (LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1 ) by A47, NAT_1:8; supposeA48: k <= LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t) ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) percases ( k = 0 or k <> 0 ) ; suppose k = 0 ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) hence IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) by A11, A22, EXTPRO_1:2; ::_thesis: verum end; suppose k <> 0 ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) then consider kn being Nat such that A49: k = kn + 1 by NAT_1:6; reconsider kn = kn as Element of NAT by ORDINAL1:def_12; reconsider lm = IC (Comput ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t,kn)) as Element of NAT ; kn < k by A49, XREAL_1:29; then kn < LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t) by A48, XXREAL_0:2; then (IC (Comput ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t,kn))) + 1 = IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,1)),kn)) by A14, A30, A25, A35, A17, A34, SCMPDS_7:16; then A50: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) = lm + 1 by A49, EXTPRO_1:4; IC (Comput ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t,kn)) in dom (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by A24, A6, SCMPDS_6:def_2; then lm < card (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by AFINSQ_1:66; then lm < (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 by COMPOS_1:55; then A51: lm + 1 <= (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 by INT_1:7; (card F3()) + 2 < (card F3()) + 4 by XREAL_1:6; then lm + 1 < (card F3()) + 4 by A20, A51, XXREAL_0:2; then lm + 1 < card (stop (for-down (F4(),F5(),F6(),F3()))) by Lm1; hence IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) by A50, AFINSQ_1:66; ::_thesis: verum end; end; end; supposeA52: k = (LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 in dom (stop (for-down (F4(),F5(),F6(),F3()))) by A26, COMPOS_1:62; hence IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) by A14, A30, A25, A35, A17, A34, A27, A52, SCMPDS_7:18; ::_thesis: verum end; end; end; end; suppose k >= ((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,b1)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) then consider nn being Nat such that A53: k = (((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k) = Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))),nn) by A53, EXTPRO_1:4 .= Comput (((Q +* (stop (for-down (F4(),F5(),F6(),F3())))) +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1))),nn) ; hence IC (Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,k)) in dom (stop (for-down (F4(),F5(),F6(),F3()))) by A46, A41, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence for-down (F4(),F5(),F6(),F3()) is_closed_on t,Q by A6, SCMPDS_6:def_2; ::_thesis: for-down (F4(),F5(),F6(),F3()) is_halting_on t,Q for-down (F4(),F5(),F6(),F3()) is_halting_on Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)),Q +* (stop (for-down (F4(),F5(),F6(),F3()))) by A5, A45, A42, A44, A13, A41; then Q +* (stop (for-down (F4(),F5(),F6(),F3()))) halts_on Comput ((Q +* (stop (for-down (F4(),F5(),F6(),F3())))),t,(((LifeSpan ((Q +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),t)) + 1) + 1)) by A43, A41, SCMPDS_6:def_3; then Q +* (stop (for-down (F4(),F5(),F6(),F3()))) halts_on t by EXTPRO_1:22; hence for-down (F4(),F5(),F6(),F3()) is_halting_on t,Q by A6, SCMPDS_6:def_3; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; A54: S1[ 0 ] by SCMPDS_7:44; A55: for k being Element of NAT holds S1[k] from NAT_1:sch_1(A54, A4); percases ( F1() . (DataLoc ((F1() . F4()),F5())) <= 0 or F1() . (DataLoc ((F1() . F4()),F5())) > 0 ) ; suppose F1() . (DataLoc ((F1() . F4()),F5())) <= 0 ; ::_thesis: ( for-down (F4(),F5(),F6(),F3()) is_closed_on F1(),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on F1(),F2() ) hence ( for-down (F4(),F5(),F6(),F3()) is_closed_on F1(),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on F1(),F2() ) by SCMPDS_7:44; ::_thesis: verum end; suppose F1() . (DataLoc ((F1() . F4()),F5())) > 0 ; ::_thesis: ( for-down (F4(),F5(),F6(),F3()) is_closed_on F1(),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on F1(),F2() ) then reconsider m = F1() . (DataLoc ((F1() . F4()),F5())) as Element of NAT by INT_1:3; S1[m] by A55; hence ( for-down (F4(),F5(),F6(),F3()) is_closed_on F1(),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on F1(),F2() ) by A2; ::_thesis: verum end; end; end; scheme :: SCPISORT:sch 2 ForDownExec{ P1[ set ], F1() -> 0 -started State of SCMPDS, F2() -> Instruction-Sequence of SCMPDS, F3() -> halt-free shiftable Program of SCMPDS, F4() -> Int_position, F5() -> Integer, F6() -> Element of NAT } : IExec ((for-down (F4(),F5(),F6(),F3())),F2(),F1()) = IExec ((for-down (F4(),F5(),F6(),F3())),F2(),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))) provided A1: F6() > 0 and A2: F1() . (DataLoc ((F1() . F4()),F5())) > 0 and A3: P1[F1()] and A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F4() = F1() . F4() & t . (DataLoc ((F1() . F4()),F5())) > 0 holds ( (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() & (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc ((F1() . F4()),F5())) = (t . (DataLoc ((F1() . F4()),F5()))) - F6() & F3() is_closed_on t,Q & F3() is_halting_on t,Q & P1[ Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t))] ) proof set i1 = (F4(),F5()) <=0_goto ((card F3()) + 3); set J = F3() ';' (AddTo (F4(),F5(),(- F6()))); set i3 = goto (- ((card F3()) + 2)); set FOR = for-down (F4(),F5(),F6(),F3()); set pFOR = stop (for-down (F4(),F5(),F6(),F3())); set P1 = F2() +* (stop (for-down (F4(),F5(),F6(),F3()))); set PJ = F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))); set mJ = LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1()); set m1 = (LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2; set s2 = Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())); set m2 = LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))); A5: stop (F3() ';' (AddTo (F4(),F5(),(- F6())))) c= F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by FUNCT_4:25; A6: Initialize F1() = F1() by MEMSTR_0:44; A7: ( F3() is_closed_on F1(),F2() & F3() is_halting_on F1(),F2() ) by A2, A3, A4; then F3() ';' (AddTo (F4(),F5(),(- F6()))) is_closed_on F1(),F2() by Th5; then A8: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_closed_on F1(),F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by A6, SCMPDS_6:24; A9: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_halting_on F1(),F2() by A7, Th5; then A10: F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) halts_on F1() by A6, SCMPDS_6:def_3; (F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))) +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) halts_on Initialize F1() by A9, SCMPDS_6:def_3; then A11: F3() ';' (AddTo (F4(),F5(),(- F6()))) is_halting_on F1(),F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6()))))) by SCMPDS_6:def_3; A12: for-down (F4(),F5(),F6(),F3()) = ((F4(),F5()) <=0_goto ((card F3()) + 3)) ';' ((F3() ';' (AddTo (F4(),F5(),(- F6())))) ';' (goto (- ((card F3()) + 2)))) by Lm2; A13: Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(0 + 1)) = Following ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),0))) by EXTPRO_1:3 .= Following ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()) by EXTPRO_1:2 .= Exec (((F4(),F5()) <=0_goto ((card F3()) + 3)),F1()) by A12, A6, SCMPDS_6:11 ; set m3 = (LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1; set s4 = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1); set P4 = F2() +* (stop (for-down (F4(),F5(),F6(),F3()))); set b = DataLoc ((F1() . F4()),F5()); A14: card (F3() ';' (AddTo (F4(),F5(),(- F6())))) = (card F3()) + 1 by SCMP_GCD:4; set s6 = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1)); set P6 = F2() +* (stop (for-down (F4(),F5(),F6(),F3()))); set s5 = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1()))); set l1 = (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1; set P5 = F2() +* (stop (for-down (F4(),F5(),F6(),F3()))); (card F3()) + 2 < (card F3()) + 3 by XREAL_1:6; then A15: (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 in dom (for-down (F4(),F5(),F6(),F3())) by A14, SCMPDS_7:42; A16: Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1)) = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1()))) by EXTPRO_1:4; for x being Int_position holds F1() . x = (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)) . x by A13, SCMPDS_2:56; then A17: DataPart F1() = DataPart (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)) by SCMPDS_4:8; A18: IC F1() = 0 by MEMSTR_0:def_12; A19: stop (for-down (F4(),F5(),F6(),F3())) c= F2() +* (stop (for-down (F4(),F5(),F6(),F3()))) by FUNCT_4:25; A20: for-down (F4(),F5(),F6(),F3()) c= stop (for-down (F4(),F5(),F6(),F3())) by AFINSQ_1:74; then A21: for-down (F4(),F5(),F6(),F3()) c= F2() +* (stop (for-down (F4(),F5(),F6(),F3()))) by A19, XBOOLE_1:1; Shift ((F3() ';' (AddTo (F4(),F5(),(- F6())))),1) c= for-down (F4(),F5(),F6(),F3()) by Lm3; then Shift ((F3() ';' (AddTo (F4(),F5(),(- F6())))),1) c= stop (for-down (F4(),F5(),F6(),F3())) by A20, XBOOLE_1:1; then A22: Shift ((F3() ';' (AddTo (F4(),F5(),(- F6())))),1) c= F2() +* (stop (for-down (F4(),F5(),F6(),F3()))) by A19, XBOOLE_1:1; set m0 = LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()); set m4 = ((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1; set s7 = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1)); A23: IC (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)) = succ (IC F1()) by A2, A13, SCMPDS_2:56 .= 0 + 1 by A18 ; then A24: IC (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())))) = (card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1 by A5, A11, A8, A17, A22, SCMPDS_7:18; then A25: CurInstr ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1)))) = (F2() +* (stop (for-down (F4(),F5(),F6(),F3())))) . ((card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1) by A16, PBOOLE:143 .= (for-down (F4(),F5(),F6(),F3())) . ((card (F3() ';' (AddTo (F4(),F5(),(- F6()))))) + 1) by A15, A21, GRFUNC_1:2 .= goto (- ((card F3()) + 2)) by A14, SCMPDS_7:43 ; A26: Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1)) = Following ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card F3()) + 2))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1)))) by A25 ; A27: DataPart (Comput ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1(),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())))) = DataPart (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())))) by A5, A11, A8, A23, A17, A22, SCMPDS_7:18; now__::_thesis:_for_x_being_Int_position_holds_(Comput_((F2()_+*_(stop_(for-down_(F4(),F5(),F6(),F3())))),F1(),(((LifeSpan_((F2()_+*_(stop_(F3()_';'_(AddTo_(F4(),F5(),(-_F6())))))),F1()))_+_1)_+_1)))_._x_=_(Initialize_(IExec_((F3()_';'_(AddTo_(F4(),F5(),(-_F6())))),F2(),F1())))_._x let x be Int_position; ::_thesis: (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1))) . x = (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . x not x in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; then A28: (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . x = (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())) . x by FUNCT_4:11; (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),1)),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())))) . x = (Comput ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1(),(LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())))) . x by A27, SCMPDS_4:8 .= (Result ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) . x by A10, EXTPRO_1:23 .= (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())) . x by SCMPDS_4:def_5 ; hence (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1))) . x = (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . x by A16, A26, A28, SCMPDS_2:54; ::_thesis: verum end; then A29: DataPart (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1))) = DataPart (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) by SCMPDS_4:8; A30: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F4() = F1() . F4() & t . (DataLoc ((F1() . F4()),F5())) > 0 holds ( (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() & (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc ((F1() . F4()),F5())) = (t . (DataLoc ((F1() . F4()),F5()))) - F6() & F3() is_closed_on t,Q & F3() is_halting_on t,Q & P1[ Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t))] ) by A4; A31: P1[F1()] by A3; ( for-down (F4(),F5(),F6(),F3()) is_closed_on F1(),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on F1(),F2() ) from SCPISORT:sch_1(A1, A31, A30); then A32: F2() +* (stop (for-down (F4(),F5(),F6(),F3()))) halts_on F1() by A6, SCMPDS_6:def_3; set Es = IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()); set bj = DataLoc (((Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . F4()),F5()); (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . F4() = (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())) . F4() by SCMPDS_5:15 .= F1() . F4() by A2, A3, A4 ; then A33: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F4() = (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . F4() & t . (DataLoc (((Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . F4()),F5())) > 0 holds ( (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() & (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc (((Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . F4()),F5())) = (t . (DataLoc (((Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) . F4()),F5()))) - F6() & F3() is_closed_on t,Q & F3() is_halting_on t,Q & P1[ Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),Q,t))] ) by A4; A34: P1[ Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))] by A2, A3, A4; ( for-down (F4(),F5(),F6(),F3()) is_closed_on Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())),F2() & for-down (F4(),F5(),F6(),F3()) is_halting_on Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())),F2() ) from SCPISORT:sch_1(A1, A34, A33); then A35: F2() +* (stop (for-down (F4(),F5(),F6(),F3()))) halts_on Initialize (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) by SCMPDS_6:def_3; IC (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1) + 1))) = ICplusConst ((Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 1))),(0 - ((card F3()) + 2))) by A26, SCMPDS_2:54 .= 0 by A14, A24, A16, SCMPDS_7:1 ; then A36: IC (Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))) = IC (Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2))) by MEMSTR_0:47; A37: Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2)) = Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())) by A29, A36, MEMSTR_0:78; then CurInstr ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2)))) = (F4(),F5()) <=0_goto ((card F3()) + 3) by A12, SCMPDS_6:11; then LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()) > (LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2 by A32, EXTPRO_1:36, SCMPDS_6:17; then consider nn being Nat such that A38: LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()) = ((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2) + (LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))))))) = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))),(LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))))) by A37, EXTPRO_1:4; then CurInstr ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2) + (LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))))))))) = halt SCMPDS by A35, EXTPRO_1:def_15; then ((LifeSpan ((F2() +* (stop (F3() ';' (AddTo (F4(),F5(),(- F6())))))),F1())) + 2) + (LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))))) >= LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()) by A32, EXTPRO_1:def_15; then A39: LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))) >= nn by A38, XREAL_1:6; A40: Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1(),(LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()))) = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))),nn) by A37, A38, EXTPRO_1:4; then CurInstr ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))),nn))) = halt SCMPDS by A32, EXTPRO_1:def_15; then nn >= LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))) by A35, EXTPRO_1:def_15; then nn = LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))) by A39, XXREAL_0:1; then Result ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),F1()) = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))),(LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))))) by A32, A40, EXTPRO_1:23; hence IExec ((for-down (F4(),F5(),F6(),F3())),F2(),F1()) = Comput ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1()))),(LifeSpan ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))))) by SCMPDS_4:def_5 .= Result ((F2() +* (stop (for-down (F4(),F5(),F6(),F3())))),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))) by A35, EXTPRO_1:23 .= IExec ((for-down (F4(),F5(),F6(),F3())),F2(),(Initialize (IExec ((F3() ';' (AddTo (F4(),F5(),(- F6())))),F2(),F1())))) by SCMPDS_4:def_5 ; ::_thesis: verum end; scheme :: SCPISORT:sch 3 ForDownEnd{ P1[ set ], F1() -> 0 -started State of SCMPDS, F2() -> halt-free shiftable Program of SCMPDS, F3() -> Instruction-Sequence of SCMPDS, F4() -> Int_position, F5() -> Integer, F6() -> Element of NAT } : ( (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1())) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1()))] ) provided A1: F6() > 0 and A2: P1[F1()] and A3: for Q being Instruction-Sequence of SCMPDS for t being 0 -started State of SCMPDS st P1[t] & t . F4() = F1() . F4() & t . (DataLoc ((F1() . F4()),F5())) > 0 holds ( (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() & (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc ((F1() . F4()),F5())) = (t . (DataLoc ((F1() . F4()),F5()))) - F6() & F2() is_closed_on t,Q & F2() is_halting_on t,Q & P1[ Initialize (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),Q,t))] ) proof set b = DataLoc ((F1() . F4()),F5()); set FR = for-down (F4(),F5(),F6(),F2()); defpred S1[ Nat] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . (DataLoc ((F1() . F4()),F5())) <= $1 & t . F4() = F1() . F4() & P1[t] holds ( (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t))] ); A4: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (DataLoc ((F1() . F4()),F5())) <= 0 & t . F4() = F1() . F4() & P1[t] holds ( (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t))] ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (DataLoc ((F1() . F4()),F5())) <= 0 & t . F4() = F1() . F4() & P1[t] implies ( (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t))] ) ) assume that A5: ( t . (DataLoc ((F1() . F4()),F5())) <= 0 & t . F4() = F1() . F4() ) and A6: P1[t] ; ::_thesis: ( (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t))] ) A7: Initialize t = t by MEMSTR_0:44; hence (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t)) . (DataLoc ((F1() . F4()),F5())) <= 0 by A5, SCMPDS_7:47; ::_thesis: P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t))] for x being Int_position holds (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t)) . x = t . x by A5, A7, SCMPDS_7:47; hence P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,t))] by A6, Th4, A7; ::_thesis: verum end; A8: now__::_thesis:_for_k_being_Element_of_NAT_st_S1[k]_holds_ S1[k_+_1] let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A9: S1[k] ; ::_thesis: S1[k + 1] thus S1[k + 1] ::_thesis: verum proof let u be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st u . (DataLoc ((F1() . F4()),F5())) <= k + 1 & u . F4() = F1() . F4() & P1[u] holds ( (IExec ((for-down (F4(),F5(),F6(),F2())),Q,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),Q,u))] ) let U be Instruction-Sequence of SCMPDS; ::_thesis: ( u . (DataLoc ((F1() . F4()),F5())) <= k + 1 & u . F4() = F1() . F4() & P1[u] implies ( (IExec ((for-down (F4(),F5(),F6(),F2())),U,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),U,u))] ) ) assume that A10: u . (DataLoc ((F1() . F4()),F5())) <= k + 1 and A11: u . F4() = F1() . F4() and A12: P1[u] ; ::_thesis: ( (IExec ((for-down (F4(),F5(),F6(),F2())),U,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),U,u))] ) percases ( u . (DataLoc ((F1() . F4()),F5())) <= 0 or u . (DataLoc ((F1() . F4()),F5())) > 0 ) ; suppose u . (DataLoc ((F1() . F4()),F5())) <= 0 ; ::_thesis: ( (IExec ((for-down (F4(),F5(),F6(),F2())),U,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),U,u))] ) hence ( (IExec ((for-down (F4(),F5(),F6(),F2())),U,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),U,u))] ) by A4, A11, A12; ::_thesis: verum end; supposeA13: u . (DataLoc ((F1() . F4()),F5())) > 0 ; ::_thesis: ( (IExec ((for-down (F4(),F5(),F6(),F2())),U,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),U,u))] ) set Ad = AddTo (F4(),F5(),(- F6())); set Iu = IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u); A14: ( (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . F4() = F1() . F4() & P1[ Initialize (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u))] ) by A3, A11, A12, A13; (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . (DataLoc ((F1() . F4()),F5())) = (u . (DataLoc ((F1() . F4()),F5()))) - F6() by A3, A11, A12, A13; then ((IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . (DataLoc ((F1() . F4()),F5()))) + 1 <= u . (DataLoc ((F1() . F4()),F5())) by A1, INT_1:7, XREAL_1:44; then ((IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . (DataLoc ((F1() . F4()),F5()))) + 1 <= k + 1 by A10, XXREAL_0:2; then A15: (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . (DataLoc ((F1() . F4()),F5())) <= k by XREAL_1:6; A16: P1[u] by A12; A17: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F4() = u . F4() & t . (DataLoc ((u . F4()),F5())) > 0 holds ( (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . F4() = t . F4() & (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),Q,t)) . (DataLoc ((u . F4()),F5())) = (t . (DataLoc ((u . F4()),F5()))) - F6() & F2() is_closed_on t,Q & F2() is_halting_on t,Q & P1[ Initialize (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),Q,t))] ) by A3, A11; A18: u . (DataLoc ((u . F4()),F5())) > 0 by A11, A13; A19: (Initialize (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u))) . (DataLoc ((F1() . F4()),F5())) = (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . (DataLoc ((F1() . F4()),F5())) by SCMPDS_5:15; A20: (Initialize (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u))) . F4() = (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)) . F4() by SCMPDS_5:15; IExec ((for-down (F4(),F5(),F6(),F2())),U,u) = IExec ((for-down (F4(),F5(),F6(),F2())),U,(Initialize (IExec ((F2() ';' (AddTo (F4(),F5(),(- F6())))),U,u)))) from SCPISORT:sch_2(A1, A18, A16, A17); hence ( (IExec ((for-down (F4(),F5(),F6(),F2())),U,u)) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),U,u))] ) by A9, A15, A14, A19, A20; ::_thesis: verum end; end; end; end; A21: for k being Element of NAT holds S1[k] from NAT_1:sch_1(A4, A8); percases ( F1() . (DataLoc ((F1() . F4()),F5())) > 0 or F1() . (DataLoc ((F1() . F4()),F5())) <= 0 ) ; suppose F1() . (DataLoc ((F1() . F4()),F5())) > 0 ; ::_thesis: ( (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1())) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1()))] ) then reconsider m = F1() . (DataLoc ((F1() . F4()),F5())) as Element of NAT by INT_1:3; S1[m] by A21; hence ( (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1())) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1()))] ) by A2; ::_thesis: verum end; suppose F1() . (DataLoc ((F1() . F4()),F5())) <= 0 ; ::_thesis: ( (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1())) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1()))] ) hence ( (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1())) . (DataLoc ((F1() . F4()),F5())) <= 0 & P1[ Initialize (IExec ((for-down (F4(),F5(),F6(),F2())),F3(),F1()))] ) by A2, A4; ::_thesis: verum end; end; end; theorem Th11: :: SCPISORT:11 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free shiftable Program of SCMPDS for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let a, x, y be Int_position; ::_thesis: for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let i, c be Integer; ::_thesis: for n being Element of NAT st n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let n be Element of NAT ; ::_thesis: ( n > 0 & s . x >= (s . y) + c & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) implies ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) ) set b = DataLoc ((s . a),i); set J = I ';' (AddTo (a,i,(- n))); assume A1: n > 0 ; ::_thesis: ( not s . x >= (s . y) + c or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) or ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) ) defpred S1[ set ] means ex t being State of SCMPDS st ( t = $1 & t . x >= (t . y) + c ); assume A2: s . x >= (s . y) + c ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) or ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) ) A3: S1[s] by A2; assume A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ; ::_thesis: ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) A5: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_S1[t]_&_t_._a_=_s_._a_&_t_._(DataLoc_((s_._a),i))_>_0_holds_ (_(IExec_((I_';'_(AddTo_(a,i,(-_n)))),Q,t))_._a_=_t_._a_&_(IExec_((I_';'_(AddTo_(a,i,(-_n)))),Q,t))_._(DataLoc_((s_._a),i))_=_(t_._(DataLoc_((s_._a),i)))_-_n_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_S1[_Initialize_(IExec_((I_';'_(AddTo_(a,i,(-_n)))),Q,t))]_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st S1[t] & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( S1[t] & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 implies ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ) ) assume that A6: S1[t] and A7: ( t . a = s . a & t . (DataLoc ((s . a),i)) > 0 ) ; ::_thesis: ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ) consider v being State of SCMPDS such that A8: v = t and A9: v . x >= (v . y) + c by A6; thus ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q ) by A4, A7, A8, A9; ::_thesis: S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] thus S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ::_thesis: verum proof take v = Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)); ::_thesis: ( v = Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) & v . x >= (v . y) + c ) thus v = Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) ; ::_thesis: v . x >= (v . y) + c v . x = (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x by SCMPDS_5:15; then v . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c by A4, A7, A8, A9; hence v . x >= (v . y) + c by SCMPDS_5:15; ::_thesis: verum end; end; ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) from SCPISORT:sch_1(A1, A3, A5); hence ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) ; ::_thesis: verum end; theorem Th12: :: SCPISORT:12 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free shiftable Program of SCMPDS for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a, x, y being Int_position for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) let a, x, y be Int_position; ::_thesis: for i, c being Integer for n being Element of NAT st n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) let i, c be Integer; ::_thesis: for n being Element of NAT st n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) holds IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) let n be Element of NAT ; ::_thesis: ( n > 0 & s . x >= (s . y) + c & s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) implies IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) ) set b = DataLoc ((s . a),i); set J = I ';' (AddTo (a,i,(- n))); assume A1: n > 0 ; ::_thesis: ( not s . x >= (s . y) + c or not s . (DataLoc ((s . a),i)) > 0 or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) or IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) ) defpred S1[ set ] means ex t being State of SCMPDS st ( t = $1 & t . x >= (t . y) + c ); assume A2: s . x >= (s . y) + c ; ::_thesis: ( not s . (DataLoc ((s . a),i)) > 0 or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) or IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) ) A3: S1[s] by A2; assume A4: s . (DataLoc ((s . a),i)) > 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ) or IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) ) assume A5: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . x >= (t . y) + c & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c ) ; ::_thesis: IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) A6: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_S1[t]_&_t_._a_=_s_._a_&_t_._(DataLoc_((s_._a),i))_>_0_holds_ (_(IExec_((I_';'_(AddTo_(a,i,(-_n)))),Q,t))_._a_=_t_._a_&_(IExec_((I_';'_(AddTo_(a,i,(-_n)))),Q,t))_._(DataLoc_((s_._a),i))_=_(t_._(DataLoc_((s_._a),i)))_-_n_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_S1[_Initialize_(IExec_((I_';'_(AddTo_(a,i,(-_n)))),Q,t))]_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st S1[t] & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 holds ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( S1[t] & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 implies ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ) ) assume that A7: S1[t] and A8: ( t . a = s . a & t . (DataLoc ((s . a),i)) > 0 ) ; ::_thesis: ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q & S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ) consider v being State of SCMPDS such that A9: v = t and A10: v . x >= (v . y) + c by A7; thus ( (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . a = t . a & (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . (DataLoc ((s . a),i)) = (t . (DataLoc ((s . a),i))) - n & I is_closed_on t,Q & I is_halting_on t,Q ) by A5, A8, A9, A10; ::_thesis: S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] thus S1[ Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t))] ::_thesis: verum proof take v = Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)); ::_thesis: ( v = Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) & v . x >= (v . y) + c ) thus v = Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) ; ::_thesis: v . x >= (v . y) + c v . x = (IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . x by SCMPDS_5:15; then v . x >= ((IExec ((I ';' (AddTo (a,i,(- n)))),Q,t)) . y) + c by A5, A8, A9, A10; hence v . x >= (v . y) + c by SCMPDS_5:15; ::_thesis: verum end; end; IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) from SCPISORT:sch_2(A1, A4, A3, A6); hence IExec ((for-down (a,i,n,I)),P,s) = IExec ((for-down (a,i,n,I)),P,(Initialize (IExec ((I ';' (AddTo (a,i,(- n)))),P,s)))) ; ::_thesis: verum end; theorem :: SCPISORT:13 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for n being Element of NAT st s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for n being Element of NAT st s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for n being Element of NAT st s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a being Int_position for i being Integer for n being Element of NAT st s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let a be Int_position; ::_thesis: for i being Integer for n being Element of NAT st s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let i be Integer; ::_thesis: for n being Element of NAT st s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) holds ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) let n be Element of NAT ; ::_thesis: ( s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ) implies ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) ) assume that A1: ( s . (DataLoc ((s . a),i)) > 0 & n > 0 & a <> DataLoc ((s . a),i) ) and A2: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q ) ; ::_thesis: ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {} holds t . x = s . x ) & t . a = s . a holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) = t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for y being Int_position st y in {} holds (IExec (I,Q,t)) . y = t . y ) ) by A2; hence ( for-down (a,i,n,I) is_closed_on s,P & for-down (a,i,n,I) is_halting_on s,P ) by A1, SCMPDS_7:48; ::_thesis: verum end; begin definition let n, p0 be Element of NAT ; func insert-sort (n,p0) -> Program of SCMPDS equals :: SCPISORT:def 2 ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)) ';' (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))); coherence ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)) ';' (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))) is Program of SCMPDS ; end; :: deftheorem defines insert-sort SCPISORT:def_2_:_ for n, p0 being Element of NAT holds insert-sort (n,p0) = ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)) ';' (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))); set j1 = AddTo (GBP,3,1); set j2 = (GBP,4) := (GBP,3); set j3 = AddTo (GBP,1,1); set j4 = (GBP,6) := (GBP,1); set k1 = (GBP,5) := ((intpos 4),(- 1)); set k2 = SubFrom (GBP,5,(intpos 4),0); set k3 = (GBP,5) := ((intpos 4),(- 1)); set k4 = ((intpos 4),(- 1)) := ((intpos 4),0); set k5 = ((intpos 4),0) := (GBP,5); set k6 = AddTo (GBP,4,(- 1)); set k7 = AddTo (GBP,6,(- 1)); set FA = Load ((GBP,6) := 0); set TR = (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1))); set IF = if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))); set B1 = (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))); set WH = while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))); set J4 = (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)); set B2 = ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))); set FR = for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))); Lm4: card ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))) = 10 proof thus card ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))) = (card (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0)))) + (card (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))) by AFINSQ_1:17 .= 2 + (card (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))) by SCMP_GCD:5 .= 2 + (((card ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1))))) + (card (Load ((GBP,6) := 0)))) + 2) by SCMPDS_6:65 .= 2 + ((((card (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1))))) + 1) + (card (Load ((GBP,6) := 0)))) + 2) by SCMP_GCD:4 .= 2 + (((((card ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5)))) + 1) + 1) + (card (Load ((GBP,6) := 0)))) + 2) by SCMP_GCD:4 .= 2 + ((((((card (((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0)))) + 1) + 1) + 1) + (card (Load ((GBP,6) := 0)))) + 2) by SCMP_GCD:4 .= 2 + (((((2 + 1) + 1) + 1) + (card (Load ((GBP,6) := 0)))) + 2) by SCMP_GCD:5 .= 2 + (((((2 + 1) + 1) + 1) + 1) + 2) by COMPOS_1:54 .= 10 ; ::_thesis: verum end; Lm5: card (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) = 16 proof thus card (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) = (card ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)))) + (card (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) by AFINSQ_1:17 .= ((card (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1)))) + 1) + (card (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) by SCMP_GCD:4 .= (((card ((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3)))) + 1) + 1) + (card (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) by SCMP_GCD:4 .= ((2 + 1) + 1) + (card (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) by SCMP_GCD:5 .= ((2 + 1) + 1) + (10 + 2) by Lm4, SCMPDS_8:17 .= 16 ; ::_thesis: verum end; set a1 = intpos 1; set a2 = intpos 2; set a3 = intpos 3; set a4 = intpos 4; set a5 = intpos 5; set a6 = intpos 6; Lm6: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . GBP = 0 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 1) = s . (intpos 1) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . GBP = 0 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 1) = s . (intpos 1) ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . GBP = 0 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 1) = s . (intpos 1) ) ) set a = GBP ; set x = DataLoc ((s . (intpos 4)),(- 1)); set y = DataLoc ((s . (intpos 4)),0); assume that A1: s . (intpos 4) >= 7 + (s . (intpos 6)) and A2: s . GBP = 0 and A3: s . (intpos 6) > 0 ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . GBP = 0 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 1) = s . (intpos 1) ) A4: 7 + (s . (intpos 6)) > 7 + 0 by A3, XREAL_1:6; set t0 = s; set t1 = Exec (((GBP,5) := ((intpos 4),(- 1))),s); set t2 = IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s); set Q2 = P; A5: DataLoc ((s . GBP),5) = intpos (0 + 5) by A2, SCMP_GCD:1; then A6: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP = 0 by A2, AMI_3:10, SCMPDS_2:47; then A7: DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP),5) = intpos (0 + 5) by SCMP_GCD:1; A8: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 4) = s . (intpos 4) by A5, AMI_3:10, SCMPDS_2:47; A9: s . (intpos 4) >= 1 + (6 + (s . (intpos 6))) by A1; then A10: (s . (intpos 4)) - 1 >= 6 + (s . (intpos 6)) by XREAL_1:19; set Fi = (GBP,6) := 0; set t02 = Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)); set Q02 = P; set t3 = IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); set t4 = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); set t5 = IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); set t6 = Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . GBP = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . GBP by SCMPDS_5:42 .= 0 by A6, A7, AMI_3:10, SCMPDS_2:50 ; then A11: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP = 0 by SCMPDS_5:15; then A12: DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5) = intpos (0 + 5) by SCMP_GCD:1; then A13: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = 0 by A11, AMI_3:10, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 4) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 4) by SCMPDS_5:42 .= s . (intpos 4) by A8, A7, AMI_3:10, SCMPDS_2:50 ; then (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 4) = s . (intpos 4) by SCMPDS_5:15; then A14: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = s . (intpos 4) by A12, AMI_3:10, SCMPDS_2:47; A15: 6 + (s . (intpos 6)) > 6 + 0 by A3, XREAL_1:6; then 0 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A10, A14, ABSVALUE:def_1; then A16: GBP <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; (s . (intpos 4)) - 1 > 0 by A3, A9, XREAL_1:19; then A17: abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) = (s . (intpos 4)) - 1 by A14, ABSVALUE:def_1; then 4 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A10, A15, XXREAL_0:2; then A18: intpos 4 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A19: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 4) by SCMPDS_5:42 .= s . (intpos 4) by A14, A18, SCMPDS_2:47 ; then 0 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, ABSVALUE:def_1; then A20: GBP <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; A21: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:42 .= 0 by A13, A16, SCMPDS_2:47 ; A22: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:41 .= 0 by A21, A20, SCMPDS_2:47 ; then A23: GBP <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by AMI_3:10, SCMP_GCD:1; A24: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 1) = s . (intpos 1) by A5, AMI_3:10, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 1) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 1) by SCMPDS_5:42 .= s . (intpos 1) by A24, A7, AMI_3:10, SCMPDS_2:50 ; then A25: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 1) = s . (intpos 1) by SCMPDS_5:15; then A26: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = s . (intpos 1) by A12, AMI_3:10, SCMPDS_2:47; A27: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:41 .= 0 by A22, A23, SCMPDS_2:48 ; then A28: DataLoc (((IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),6) = intpos (0 + 6) by SCMP_GCD:1; A29: DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),6) = intpos (0 + 6) by A11, SCMP_GCD:1; now__::_thesis:_(IExec_((if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._GBP_=_0 percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = 0 hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP by SCMPDS_5:40 .= 0 by A11, A29, AMI_3:10, SCMPDS_2:46 ; ::_thesis: verum end; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = 0 hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP by SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:41 .= 0 by A27, A28, AMI_3:10, SCMPDS_2:48 ; ::_thesis: verum end; end; end; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . GBP = 0 by SCMPDS_5:35; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 1) = s . (intpos 1) A30: intpos 1 <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by A22, AMI_3:10, SCMP_GCD:1; abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) = s . (intpos 4) by A1, A4, A19, ABSVALUE:def_1; then 1 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, XXREAL_0:2; then A31: intpos 1 <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; 1 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A10, A15, A17, XXREAL_0:2; then A32: intpos 1 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A33: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 1) by SCMPDS_5:42 .= s . (intpos 1) by A26, A32, SCMPDS_2:47 ; A34: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 1) by SCMPDS_5:41 .= s . (intpos 1) by A33, A31, SCMPDS_2:47 ; A35: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 1) by SCMPDS_5:41 .= s . (intpos 1) by A34, A30, SCMPDS_2:48 ; now__::_thesis:_(IExec_((if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_1)_=_s_._(intpos_1) percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = s . (intpos 1) hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) by SCMPDS_5:40 .= s . (intpos 1) by A25, A29, AMI_3:10, SCMPDS_2:46 ; ::_thesis: verum end; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = s . (intpos 1) hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 1) by SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 1) by SCMPDS_5:41 .= s . (intpos 1) by A35, A28, AMI_3:10, SCMPDS_2:48 ; ::_thesis: verum end; end; end; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 1) = s . (intpos 1) by SCMPDS_5:35; ::_thesis: verum end; Lm7: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 3) = s . (intpos 3) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 3) = s . (intpos 3) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 3) = s . (intpos 3) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) ) set a = GBP ; set x = DataLoc ((s . (intpos 4)),(- 1)); set y = DataLoc ((s . (intpos 4)),0); assume that A1: s . (intpos 4) >= 7 + (s . (intpos 6)) and A2: s . GBP = 0 and A3: s . (intpos 6) > 0 ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 3) = s . (intpos 3) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) set t0 = s; set t1 = Exec (((GBP,5) := ((intpos 4),(- 1))),s); set t2 = IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s); set Q2 = P; A4: 7 + (s . (intpos 6)) > 7 + 0 by A3, XREAL_1:6; then A5: abs (s . (intpos 4)) = s . (intpos 4) by A1, ABSVALUE:def_1; set Fi = (GBP,6) := 0; set t02 = Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)); set Q02 = P; set t3 = IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); set t4 = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); set t5 = IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); set t6 = Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)))); A6: DataLoc ((s . GBP),5) = intpos (0 + 5) by A2, SCMP_GCD:1; then A7: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP = 0 by A2, AMI_3:10, SCMPDS_2:47; then A8: DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP),5) = intpos (0 + 5) by SCMP_GCD:1; then A9: abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP) + 5) = 0 + 5 by XTUPLE_0:1; then abs ((s . (intpos 4)) + 0) <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP) + 5) by A1, A4, A5, XXREAL_0:2; then A10: DataLoc ((s . (intpos 4)),0) <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP),5) by XTUPLE_0:1; A11: abs ((s . GBP) + 5) = 0 + 5 by A6, XTUPLE_0:1; then abs ((s . (intpos 4)) + 0) <> abs ((s . GBP) + 5) by A1, A4, A5, XXREAL_0:2; then ( s . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & DataLoc ((s . (intpos 4)),0) <> DataLoc ((s . GBP),5) ) by XTUPLE_0:1; then A12: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_2:47; A13: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 5) = s . (DataLoc ((s . (intpos 4)),(- 1))) by A6, SCMPDS_2:47; A14: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 4) = s . (intpos 4) by A6, AMI_3:10, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:42 .= s . (DataLoc ((s . (intpos 4)),0)) by A12, A10, SCMPDS_2:50 ; then A15: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:15; A16: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP = (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . GBP by SCMPDS_5:15 .= (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . GBP by SCMPDS_5:42 .= 0 by A7, A8, AMI_3:10, SCMPDS_2:50 ; A17: DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5) = intpos (0 + 5) by A16, SCMP_GCD:1; then A18: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = 0 by A16, AMI_3:10, SCMPDS_2:47; abs (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP) + 5) = 0 + 5 by A17, XTUPLE_0:1; then abs ((s . (intpos 4)) + 0) <> abs (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP) + 5) by A1, A4, A5, XXREAL_0:2; then DataLoc ((s . (intpos 4)),0) <> DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5) by XTUPLE_0:1; then A19: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) by A15, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 4) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 4) by SCMPDS_5:42 .= s . (intpos 4) by A14, A8, AMI_3:10, SCMPDS_2:50 ; then A20: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 4) = s . (intpos 4) by SCMPDS_5:15; then A21: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = s . (intpos 4) by A17, AMI_3:10, SCMPDS_2:47; then A22: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1))) by SCMPDS_5:42 .= s . (DataLoc ((s . (intpos 4)),0)) by A21, A19, SCMPDS_2:47 ; A23: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 3) = s . (intpos 3) by A6, AMI_3:10, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 3) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 3) by SCMPDS_5:42 .= s . (intpos 3) by A23, A8, AMI_3:10, SCMPDS_2:50 ; then A24: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 3) = s . (intpos 3) by SCMPDS_5:15; then A25: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = s . (intpos 3) by A17, AMI_3:10, SCMPDS_2:47; A26: s . (intpos 4) >= 1 + (6 + (s . (intpos 6))) by A1; then A27: (s . (intpos 4)) - 1 >= 6 + (s . (intpos 6)) by XREAL_1:19; A28: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 2) = s . (intpos 2) by A6, AMI_3:10, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 2) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 2) by SCMPDS_5:42 .= s . (intpos 2) by A28, A8, AMI_3:10, SCMPDS_2:50 ; then A29: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 2) = s . (intpos 2) by SCMPDS_5:15; then A30: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = s . (intpos 2) by A17, AMI_3:10, SCMPDS_2:47; A31: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos 6) = s . (intpos 6) by A6, AMI_3:10, SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 6) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 6) by SCMPDS_5:42 .= s . (intpos 6) by A31, A8, AMI_3:10, SCMPDS_2:50 ; then (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 6) = s . (intpos 6) by SCMPDS_5:15; then A32: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) = s . (intpos 6) by A17, AMI_3:10, SCMPDS_2:47; A33: DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),6) = intpos (0 + 6) by A16, SCMP_GCD:1; A34: now__::_thesis:_(_(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))_._(DataLoc_(((Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))_._GBP),5))_<=_0_implies_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),P,s))_._(intpos_6)_=_0_) assume (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) by SCMPDS_5:40 .= 0 by A33, SCMPDS_2:46 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 by SCMPDS_5:35; ::_thesis: verum end; A35: 6 + (s . (intpos 6)) > 6 + 0 by A3, XREAL_1:6; then 0 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A27, A21, ABSVALUE:def_1; then A36: GBP <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A37: (s . (intpos 4)) - 1 > 0 by A3, A26, XREAL_1:19; then A38: abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) = (s . (intpos 4)) - 1 by A21, ABSVALUE:def_1; then 4 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A27, A35, XXREAL_0:2; then A39: intpos 4 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A40: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 4) by SCMPDS_5:42 .= s . (intpos 4) by A21, A39, SCMPDS_2:47 ; then A41: abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) = s . (intpos 4) by A1, A4, ABSVALUE:def_1; then 4 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, XXREAL_0:2; then A42: intpos 4 <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; 3 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A27, A35, A38, XXREAL_0:2; then A43: intpos 3 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; 3 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, A41, XXREAL_0:2; then A44: intpos 3 <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; A45: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 3) by SCMPDS_5:42 .= s . (intpos 3) by A25, A43, SCMPDS_2:47 ; A46: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 3) by SCMPDS_5:41 .= s . (intpos 3) by A45, A44, SCMPDS_2:47 ; 2 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, A41, XXREAL_0:2; then A47: intpos 2 <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; A48: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 4) by SCMPDS_5:41 .= s . (intpos 4) by A40, A42, SCMPDS_2:47 ; A49: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:42 .= 0 by A18, A36, SCMPDS_2:47 ; A50: (2 * (abs ((s . (intpos 4)) + (- 1)))) + 1 = (2 * ((s . (intpos 4)) - 1)) + 1 by A27, A35, ABSVALUE:def_1; then abs ((s . (intpos 4)) + (- 1)) <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP) + 5) by A3, A27, A9, XREAL_1:6; then A51: DataLoc ((s . (intpos 4)),(- 1)) <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . GBP),5) by XTUPLE_0:1; 0 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, A40, ABSVALUE:def_1; then A52: GBP <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; A53: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:41 .= 0 by A49, A52, SCMPDS_2:47 ; then A54: GBP <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by AMI_3:10, SCMP_GCD:1; 2 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A27, A35, A38, XXREAL_0:2; then A55: intpos 2 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A56: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 2) by SCMPDS_5:42 .= s . (intpos 2) by A30, A55, SCMPDS_2:47 ; A57: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 2) by SCMPDS_5:41 .= s . (intpos 2) by A56, A47, SCMPDS_2:47 ; A58: intpos 2 <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by A53, AMI_3:10, SCMP_GCD:1; (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . GBP by SCMPDS_5:41 .= 0 by A53, A54, SCMPDS_2:48 ; then A59: DataLoc (((IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),6) = intpos (0 + 6) by SCMP_GCD:1; then A60: abs (((IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP) + 6) = 0 + 6 by XTUPLE_0:1; A61: DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) = intpos (0 + 4) by A53, SCMP_GCD:1; then A62: abs (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP) + 4) = 0 + 4 by XTUPLE_0:1; then abs ((s . (intpos 4)) + (- 1)) <> abs (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP) + 4) by A27, A35, A50, XXREAL_0:2; then A63: DataLoc ((s . (intpos 4)),(- 1)) <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by XTUPLE_0:1; A64: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 2) by SCMPDS_5:41 .= s . (intpos 2) by A57, A58, SCMPDS_2:48 ; now__::_thesis:_(IExec_((if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_2)_=_s_._(intpos_2) percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = s . (intpos 2) hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) by SCMPDS_5:40 .= s . (intpos 2) by A29, A33, AMI_3:10, SCMPDS_2:46 ; ::_thesis: verum end; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = s . (intpos 2) hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 2) by SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 2) by SCMPDS_5:41 .= s . (intpos 2) by A64, A59, AMI_3:10, SCMPDS_2:48 ; ::_thesis: verum end; end; end; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 2) = s . (intpos 2) by SCMPDS_5:35; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 3) = s . (intpos 3) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) A65: intpos 3 <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by A53, AMI_3:10, SCMP_GCD:1; A66: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 3) by SCMPDS_5:41 .= s . (intpos 3) by A46, A65, SCMPDS_2:48 ; now__::_thesis:_(IExec_((if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_3)_=_s_._(intpos_3) percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = s . (intpos 3) hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) by SCMPDS_5:40 .= s . (intpos 3) by A24, A33, AMI_3:10, SCMPDS_2:46 ; ::_thesis: verum end; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = s . (intpos 3) hence (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 3) by SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 3) by SCMPDS_5:41 .= s . (intpos 3) by A66, A59, AMI_3:10, SCMPDS_2:48 ; ::_thesis: verum end; end; end; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 3) = s . (intpos 3) by SCMPDS_5:35; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) A67: intpos 6 <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by A53, AMI_3:10, SCMP_GCD:1; 6 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A26, A35, A38, XREAL_1:19; then A68: intpos 6 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A69: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 6) by SCMPDS_5:42 .= s . (intpos 6) by A32, A68, SCMPDS_2:47 ; 6 <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A1, A4, A41, XXREAL_0:2; then A70: intpos 6 <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; A71: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 6) by SCMPDS_5:41 .= s . (intpos 6) by A69, A70, SCMPDS_2:47 ; A72: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 6) by SCMPDS_5:41 .= s . (intpos 6) by A71, A67, SCMPDS_2:48 ; A73: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 4) by SCMPDS_5:41 .= ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1) by A61, SCMPDS_2:48 .= (s . (intpos 4)) - 1 by A48 ; A74: now__::_thesis:_(_(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))_._(DataLoc_(((Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))_._GBP),5))_>_0_implies_(_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),P,s))_._(intpos_6)_=_(s_._(intpos_6))_-_1_&_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),P,s))_._(intpos_4)_=_(s_._(intpos_4))_-_1_)_) assume A75: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 6) by SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 6) by SCMPDS_5:41 .= (s . (intpos 6)) + (- 1) by A72, A59, SCMPDS_2:48 .= (s . (intpos 6)) - 1 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 by SCMPDS_5:35; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) by A75, SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 4) by SCMPDS_5:41 .= (s . (intpos 4)) - 1 by A73, A59, AMI_3:10, SCMPDS_2:48 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 by SCMPDS_5:35; ::_thesis: verum end; hereby ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) & ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) by A3, A34; ::_thesis: verum end; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) < s . (intpos 6) by A74, XREAL_1:146; ::_thesis: verum end; end; end; hereby ::_thesis: ( ( for i being Nat st i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; supposeA76: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4) by SCMPDS_5:40 .= s . (intpos 4) by A20, A33, AMI_3:10, SCMPDS_2:46 ; then (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = s . (intpos 4) by SCMPDS_5:35; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) by A1, A4, A34, A76, XXREAL_0:2; ::_thesis: verum end; supposeA77: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) (s . (intpos 4)) - 1 >= (7 + (s . (intpos 6))) - 1 by A1, XREAL_1:9; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6)) by A74, A77; ::_thesis: verum end; end; end; A78: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_&_i_<>_(s_._(intpos_4))_-_1_&_i_<>_s_._(intpos_4)_holds_ (Exec_(((GBP,5)_:=_((intpos_4),(-_1))),s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos i) = s . (intpos i) ) assume that A79: i >= 7 and i <> (s . (intpos 4)) - 1 and i <> s . (intpos 4) ; ::_thesis: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos i) = s . (intpos i) i > 5 by A79, XXREAL_0:2; hence (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos i) = s . (intpos i) by A6, AMI_3:10, SCMPDS_2:47; ::_thesis: verum end; A80: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_&_i_<>_(s_._(intpos_4))_-_1_&_i_<>_s_._(intpos_4)_holds_ (IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos i) = s . (intpos i) ) assume that A81: i >= 7 and A82: ( i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) ) ; ::_thesis: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos i) = s . (intpos i) A83: i > 5 by A81, XXREAL_0:2; thus (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos i) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos i) by SCMPDS_5:42 .= (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (intpos i) by A8, A83, AMI_3:10, SCMPDS_2:50 .= s . (intpos i) by A78, A81, A82 ; ::_thesis: verum end; A84: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_&_i_<>_(s_._(intpos_4))_-_1_&_i_<>_s_._(intpos_4)_holds_ (Exec_(((GBP,5)_:=_((intpos_4),(-_1))),(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) ) assume that A85: i >= 7 and A86: ( i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) ) ; ::_thesis: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) i > 5 by A85, XXREAL_0:2; hence (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos i) by A17, AMI_3:10, SCMPDS_2:47 .= (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos i) by SCMPDS_5:15 .= s . (intpos i) by A80, A85, A86 ; ::_thesis: verum end; A87: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_&_i_<>_(s_._(intpos_4))_-_1_&_i_<>_s_._(intpos_4)_holds_ (IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) ) assume that A88: i >= 7 and A89: i <> (s . (intpos 4)) - 1 and A90: i <> s . (intpos 4) ; ::_thesis: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) A91: intpos i <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) proof assume intpos i = DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) ; ::_thesis: contradiction then i = abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by XTUPLE_0:1; hence contradiction by A37, A21, A89, ABSVALUE:def_1; ::_thesis: verum end; thus (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos i) by SCMPDS_5:42 .= (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by A91, SCMPDS_2:47 .= s . (intpos i) by A84, A88, A89, A90 ; ::_thesis: verum end; A92: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_&_i_<>_(s_._(intpos_4))_-_1_&_i_<>_s_._(intpos_4)_holds_ (IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) ) assume that A93: ( i >= 7 & i <> (s . (intpos 4)) - 1 ) and A94: i <> s . (intpos 4) ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) A95: intpos i <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) proof assume intpos i = DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) ; ::_thesis: contradiction then i = abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by XTUPLE_0:1; hence contradiction by A1, A4, A40, A94, ABSVALUE:def_1; ::_thesis: verum end; thus (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos i) by SCMPDS_5:41 .= (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by A95, SCMPDS_2:47 .= s . (intpos i) by A87, A93, A94 ; ::_thesis: verum end; A96: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_&_i_<>_(s_._(intpos_4))_-_1_&_i_<>_s_._(intpos_4)_holds_ (IExec_((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1)))),P,(Initialize_(IExec_((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0))),P,s)))))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) ) assume that A97: i >= 7 and A98: ( i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) ) ; ::_thesis: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = s . (intpos i) i > 4 by A97, XXREAL_0:2; then A99: intpos i <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by A53, AMI_3:10, SCMP_GCD:1; thus (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos i) by SCMPDS_5:41 .= (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by A99, SCMPDS_2:48 .= s . (intpos i) by A92, A97, A98 ; ::_thesis: verum end; hereby ::_thesis: ( ( s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) ) & ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) ) let i be Nat; ::_thesis: ( i >= 7 & i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) implies (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos b1) = s . (intpos b1) ) A100: i in NAT by ORDINAL1:def_12; set xi = intpos i; assume that A101: i >= 7 and A102: ( i <> (s . (intpos 4)) - 1 & i <> s . (intpos 4) ) ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos b1) = s . (intpos b1) A103: i > 6 by A101, XXREAL_0:2; percases ( (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 or (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ) ; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) <= 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos b1) = s . (intpos b1) then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by SCMPDS_5:40 .= (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos i) by A33, A103, AMI_3:10, SCMPDS_2:46 .= (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos i) by SCMPDS_5:15 .= s . (intpos i) by A80, A100, A101, A102 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) by SCMPDS_5:35; ::_thesis: verum end; suppose (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) > 0 ; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos b1) = s . (intpos b1) then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos i) by SCMPDS_5:41 .= (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos i) by A59, A103, AMI_3:10, SCMPDS_2:48 .= s . (intpos i) by A96, A100, A101, A102 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos i) = s . (intpos i) by SCMPDS_5:35; ::_thesis: verum end; end; end; A104: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (intpos 5) = (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (intpos 5) by SCMPDS_5:15 .= (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (intpos 5) by SCMPDS_5:42 .= (s . (DataLoc ((s . (intpos 4)),(- 1)))) - (s . (DataLoc ((s . (intpos 4)),0))) by A14, A13, A12, A8, SCMPDS_2:50 ; then A105: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),5)) = (s . (DataLoc ((s . (intpos 4)),(- 1)))) - (s . (DataLoc ((s . (intpos 4)),0))) by A16, SCMP_GCD:1; abs ((s . (intpos 4)) + (- 1)) <> abs ((s . GBP) + 5) by A3, A27, A50, A11, XREAL_1:6; then DataLoc ((s . (intpos 4)),(- 1)) <> DataLoc ((s . GBP),5) by XTUPLE_0:1; then A106: (Exec (((GBP,5) := ((intpos 4),(- 1))),s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_2:47; (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = (Exec ((SubFrom (GBP,5,(intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),s)))) . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:42 .= s . (DataLoc ((s . (intpos 4)),(- 1))) by A106, A51, SCMPDS_2:50 ; then A107: (Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:15; then A108: (Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 5) = s . (DataLoc ((s . (intpos 4)),(- 1))) by A20, A17, SCMPDS_2:47; 5 <> abs (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + (- 1)) by A27, A35, A38, XXREAL_0:2; then A109: intpos 5 <> DataLoc (((Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),(- 1)) by XTUPLE_0:1; A110: (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 5) = (Exec ((((intpos 4),(- 1)) := ((intpos 4),0)),(Exec (((GBP,5) := ((intpos 4),(- 1))),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (intpos 5) by SCMPDS_5:42 .= s . (DataLoc ((s . (intpos 4)),(- 1))) by A108, A109, SCMPDS_2:47 ; abs ((s . (intpos 4)) + 0) <> abs (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP) + 4) by A1, A4, A5, A62, XXREAL_0:2; then A111: DataLoc ((s . (intpos 4)),0) <> DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),4) by XTUPLE_0:1; A112: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0)) by A40, SCMPDS_5:41 .= (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),5)) by SCMPDS_2:47 .= s . (DataLoc ((s . (intpos 4)),(- 1))) by A49, A110, SCMP_GCD:1 ; A113: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:41 .= s . (DataLoc ((s . (intpos 4)),(- 1))) by A112, A111, SCMPDS_2:48 ; abs ((s . (intpos 4)) + (- 1)) <> abs (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)) + 0) by A50, A41; then A114: DataLoc ((s . (intpos 4)),(- 1)) <> DataLoc (((IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (intpos 4)),0) by XTUPLE_0:1; A115: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) = (Exec ((((intpos 4),0) := (GBP,5)),(IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:41 .= s . (DataLoc ((s . (intpos 4)),0)) by A22, A114, SCMPDS_2:47 ; A116: (IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) = (Exec ((AddTo (GBP,4,(- 1))),(IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:41 .= s . (DataLoc ((s . (intpos 4)),0)) by A115, A63, SCMPDS_2:48 ; hereby ::_thesis: ( s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) ) A117: DataLoc ((s . (intpos 4)),(- 1)) <> DataLoc (((IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),6) by A27, A35, A50, A59, XTUPLE_0:1; assume s . (DataLoc ((s . (intpos 4)),(- 1))) > s . (DataLoc ((s . (intpos 4)),0)) ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) then A118: (s . (DataLoc ((s . (intpos 4)),(- 1)))) - (s . (DataLoc ((s . (intpos 4)),0))) > (s . (DataLoc ((s . (intpos 4)),0))) - (s . (DataLoc ((s . (intpos 4)),0))) by XREAL_1:9; then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) by A105, SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:41 .= s . (DataLoc ((s . (intpos 4)),0)) by A116, A117, SCMPDS_2:48 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:35; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) abs ((s . (intpos 4)) + 0) <> abs (((IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP) + 6) by A1, A4, A5, A60, XXREAL_0:2; then A119: DataLoc ((s . (intpos 4)),0) <> DataLoc (((IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . GBP),6) by XTUPLE_0:1; (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) = (IExec (((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) by A105, A118, SCMPDS_6:73 .= (Exec ((AddTo (GBP,6,(- 1))),(IExec ((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))))) . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:41 .= s . (DataLoc ((s . (intpos 4)),(- 1))) by A113, A119, SCMPDS_2:48 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:35; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) thus ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = (s . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 4) = (s . (intpos 4)) - 1 ) by A16, A104, A74, A118, SCMP_GCD:1; ::_thesis: verum end; A120: abs (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP) + 6) = 0 + 6 by A33, XTUPLE_0:1; hereby ::_thesis: verum A121: DataLoc ((s . (intpos 4)),(- 1)) <> DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),6) by A27, A35, A50, A33, XTUPLE_0:1; assume s . (DataLoc ((s . (intpos 4)),(- 1))) <= s . (DataLoc ((s . (intpos 4)),0)) ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) then A122: (s . (DataLoc ((s . (intpos 4)),(- 1)))) - (s . (DataLoc ((s . (intpos 4)),0))) <= (s . (DataLoc ((s . (intpos 4)),0))) - (s . (DataLoc ((s . (intpos 4)),0))) by XREAL_1:9; then (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) by A105, SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:40 .= s . (DataLoc ((s . (intpos 4)),(- 1))) by A107, A121, SCMPDS_2:46 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),(- 1))) = s . (DataLoc ((s . (intpos 4)),(- 1))) by SCMPDS_5:35; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 ) abs ((s . (intpos 4)) + 0) <> abs (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP) + 6) by A1, A4, A5, A120, XXREAL_0:2; then A123: DataLoc ((s . (intpos 4)),0) <> DataLoc (((Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))) . GBP),6) by XTUPLE_0:1; (IExec ((if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) = (IExec ((Load ((GBP,6) := 0)),P,(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) by A105, A122, SCMPDS_6:74 .= (Exec (((GBP,6) := 0),(Initialize (IExec ((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))),P,s))))) . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:40 .= s . (DataLoc ((s . (intpos 4)),0)) by A15, A123, SCMPDS_2:46 ; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (DataLoc ((s . (intpos 4)),0)) = s . (DataLoc ((s . (intpos 4)),0)) by SCMPDS_5:35; ::_thesis: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 thus (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)) . (intpos 6) = 0 by A16, A104, A34, A122, SCMP_GCD:1; ::_thesis: verum end; end; Lm8: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) & s . GBP = 0 holds ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on s,P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) & s . GBP = 0 holds ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on s,P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) & s . GBP = 0 implies ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on s,P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on s,P ) ) set a = GBP ; set b = DataLoc ((s . GBP),6); assume that A1: s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) and A2: s . GBP = 0 ; ::_thesis: ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on s,P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on s,P ) A3: DataLoc ((s . GBP),6) = intpos (0 + 6) by A2, SCMP_GCD:1; A4: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_(_for_x_being_Int_position_st_x_in_{(intpos_4)}_holds_ t_._x_>=_7_+_(t_._(DataLoc_((s_._GBP),6)))_)_&_t_._GBP_=_s_._GBP_&_t_._(DataLoc_((s_._GBP),6))_>_0_holds_ (_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._GBP_=_t_._GBP_&_(((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))_is_closed_on_t,Q_&_(((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))_is_halting_on_t,Q_&_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._(DataLoc_((s_._GBP),6))_<_t_._(DataLoc_((s_._GBP),6))_&_(_for_x_being_Int_position_st_x_in_{(intpos_4)}_holds_ (IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._x_>=_7_+_((IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._(DataLoc_((s_._GBP),6)))_)_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((s . GBP),6))) ) & t . GBP = s . GBP & t . (DataLoc ((s . GBP),6)) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( ( for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((s . GBP),6))) ) & t . GBP = s . GBP & t . (DataLoc ((s . GBP),6)) > 0 implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) ) assume that A5: for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((s . GBP),6))) and A6: ( t . GBP = s . GBP & t . (DataLoc ((s . GBP),6)) > 0 ) ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) set Bt = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t); A7: intpos 4 in {(intpos 4)} by TARSKI:def_1; then A8: t . (intpos 4) >= 7 + (t . (intpos 6)) by A3, A5; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP by A2, A3, A6, Lm6; ::_thesis: ( (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) thus ( (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q ) by SCMPDS_6:20, SCMPDS_6:21; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) thus (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) by A2, A3, A6, A8, Lm7; ::_thesis: for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) t . (intpos 4) >= 7 + (t . (DataLoc ((s . GBP),6))) by A5, A7; then (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6)) by A2, A3, A6, Lm7; hence for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) by A3, TARSKI:def_1; ::_thesis: verum end; for x being Int_position st x in {(intpos 4)} holds s . x >= 7 + (s . (DataLoc ((s . GBP),6))) by A1, TARSKI:def_1; hence ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on s,P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on s,P ) by A4, SCMPDS_8:29; ::_thesis: verum end; Lm9: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) & s . GBP = 0 holds ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) & s . GBP = 0 holds ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) & s . GBP = 0 implies ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) ) set b = DataLoc ((s . GBP),6); set a = GBP ; assume that A1: s . (intpos 4) >= 7 + (s . (DataLoc ((s . GBP),6))) and A2: s . GBP = 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) A3: DataLoc ((s . GBP),6) = intpos (0 + 6) by A2, SCMP_GCD:1; defpred S1[ Nat] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . (intpos 6) <= $1 & t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 holds ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ); A4: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (intpos 6) <= 0 & t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 holds ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (intpos 6) <= 0 & t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 implies ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) ) assume that A5: t . (intpos 6) <= 0 and t . (intpos 4) >= 7 + (t . (intpos 6)) and A6: t . GBP = 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) A7: DataLoc ((t . GBP),6) = intpos (0 + 6) by A6, SCMP_GCD:1; hence (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 by A5, A6, SCMPDS_8:23; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) thus ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) by A5, A7, SCMPDS_8:23; ::_thesis: verum end; A8: for k being Element of NAT st S1[k] holds S1[k + 1] proof let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A9: S1[k] ; ::_thesis: S1[k + 1] thus S1[k + 1] ::_thesis: verum proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (intpos 6) <= k + 1 & t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 holds ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (intpos 6) <= k + 1 & t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 implies ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) ) set bt = DataLoc ((t . GBP),6); assume that A10: t . (intpos 6) <= k + 1 and A11: t . (intpos 4) >= 7 + (t . (intpos 6)) and A12: t . GBP = 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) A13: DataLoc ((t . GBP),6) = intpos (0 + 6) by A12, SCMP_GCD:1; percases ( t . (DataLoc ((t . GBP),6)) <= 0 or t . (DataLoc ((t . GBP),6)) > 0 ) ; suppose t . (DataLoc ((t . GBP),6)) <= 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) hence ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) by A12, SCMPDS_8:23; ::_thesis: verum end; supposeA14: t . (DataLoc ((t . GBP),6)) > 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) A15: for v being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {(intpos 4)} holds v . x >= 7 + (v . (DataLoc ((t . GBP),6))) ) & v . GBP = t . GBP & v . (DataLoc ((t . GBP),6)) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . GBP = v . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on v,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on v,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . (DataLoc ((t . GBP),6))) ) ) proof let v be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {(intpos 4)} holds v . x >= 7 + (v . (DataLoc ((t . GBP),6))) ) & v . GBP = t . GBP & v . (DataLoc ((t . GBP),6)) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . GBP = v . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on v,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on v,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,v)) . (DataLoc ((t . GBP),6))) ) ) let V be Instruction-Sequence of SCMPDS; ::_thesis: ( ( for x being Int_position st x in {(intpos 4)} holds v . x >= 7 + (v . (DataLoc ((t . GBP),6))) ) & v . GBP = t . GBP & v . (DataLoc ((t . GBP),6)) > 0 implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . GBP = v . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on v,V & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on v,V & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6))) ) ) ) A16: Initialize v = v by MEMSTR_0:44; assume that A17: for x being Int_position st x in {(intpos 4)} holds v . x >= 7 + (v . (DataLoc ((t . GBP),6))) and A18: ( v . GBP = t . GBP & v . (DataLoc ((t . GBP),6)) > 0 ) ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . GBP = v . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on v,V & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on v,V & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6))) ) ) set Iv = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,(Initialize (Initialize v))); A19: v . GBP = (Initialize v) . GBP by SCMPDS_5:15; A20: v . (intpos 4) = (Initialize v) . (intpos 4) by SCMPDS_5:15; A21: v . (intpos 6) = (Initialize v) . (intpos 6) by SCMPDS_5:15; A22: intpos 4 in {(intpos 4)} by TARSKI:def_1; then A23: v . (intpos 4) >= 7 + (v . (intpos 6)) by A13, A17; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . GBP = v . GBP by A12, A13, A18, Lm6; ::_thesis: ( (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on v,V & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on v,V & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6))) ) ) thus ( (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on v,V & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on v,V ) by SCMPDS_6:20, SCMPDS_6:21; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6))) ) ) thus (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6)) < v . (DataLoc ((t . GBP),6)) by A12, A13, A18, A23, Lm7; ::_thesis: for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6))) v . (intpos 4) >= 7 + (v . (DataLoc ((t . GBP),6))) by A17, A22; then (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,(Initialize (Initialize v)))) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,(Initialize (Initialize v)))) . (intpos 6)) by A12, A13, A18, Lm7, A19, A20, A21; hence for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),V,v)) . (DataLoc ((t . GBP),6))) by A13, A16, TARSKI:def_1; ::_thesis: verum end; set It = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t); set IT = Q; A24: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 1) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 1) by SCMPDS_5:15; A25: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 6) by SCMPDS_5:15; A26: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 4) by SCMPDS_5:15; A27: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . GBP by SCMPDS_5:15; A28: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 2) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 2) by SCMPDS_5:15; A29: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 3) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 3) by SCMPDS_5:15; (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6) < t . (intpos 6) by A11, A12, A13, A14, Lm7; then ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6)) + 1 <= t . (intpos 6) by INT_1:7; then ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6)) + 1 <= k + 1 by A10, XXREAL_0:2; then A30: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6) <= k by XREAL_1:6; A31: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = 0 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6)) ) by A11, A12, A13, A14, Lm6, Lm7; then A32: (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos 1) = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 1) by A9, A30, A24, A25, A26, A27; A33: (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos 3) = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 3) by A9, A31, A30, A25, A26, A27, A29; A34: (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos 2) = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 2) by A9, A31, A30, A25, A26, A27, A28; A35: for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((t . GBP),6))) by A11, A13, TARSKI:def_1; (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . GBP = 0 by A9, A31, A30, A25, A26, A27; hence (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . GBP = 0 by A14, A35, A15, SCMPDS_8:29; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 1) = t . (intpos 1) by A11, A12, A13, A14, Lm6; hence (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 1) = t . (intpos 1) by A14, A35, A15, A32, SCMPDS_8:29; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) ) (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 2) = t . (intpos 2) by A11, A12, A13, A14, Lm7; hence (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 2) = t . (intpos 2) by A14, A35, A15, A34, SCMPDS_8:29; ::_thesis: (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 3) = t . (intpos 3) by A11, A12, A13, A14, Lm7; hence (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos 3) = t . (intpos 3) by A14, A35, A15, A33, SCMPDS_8:29; ::_thesis: verum end; end; end; end; percases ( s . (intpos 6) <= 0 or s . (intpos 6) > 0 ) ; suppose s . (intpos 6) <= 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) hence ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) by A2, A3, SCMPDS_8:23; ::_thesis: verum end; suppose s . (intpos 6) > 0 ; ::_thesis: ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) then reconsider m = s . (intpos 6) as Element of NAT by INT_1:3; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A4, A8); then S1[m] ; hence ( (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . GBP = 0 & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 1) = s . (intpos 1) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s)) . (intpos 3) = s . (intpos 3) ) by A1, A2, A3; ::_thesis: verum end; end; end; Lm10: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . GBP = 0 holds ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = 0 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . GBP = 0 holds ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = 0 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: ( s . GBP = 0 implies ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = 0 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) ) A1: Initialize s = s by MEMSTR_0:44; set t0 = Initialize s; set t1 = IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s); set t2 = IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s); set t3 = IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s); set t4 = Exec ((AddTo (GBP,3,1)),s); assume A2: s . GBP = 0 ; ::_thesis: ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = 0 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) then A3: (Initialize s) . GBP = 0 by SCMPDS_5:15; then A4: DataLoc (((Initialize s) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; then A5: (Exec ((AddTo (GBP,3,1)),s)) . GBP = 0 by A3, A2, AMI_3:10, SCMPDS_2:48; then A6: DataLoc (((Exec ((AddTo (GBP,3,1)),s)) . GBP),4) = intpos (0 + 4) by SCMP_GCD:1; A7: (Exec ((AddTo (GBP,3,1)),s)) . (intpos 2) = s . (intpos 2) by A4, A2, A3, AMI_3:10, SCMPDS_2:48; A8: (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos 2) = (Exec (((GBP,4) := (GBP,3)),(Exec ((AddTo (GBP,3,1)),s)))) . (intpos 2) by SCMPDS_5:42 .= s . (intpos 2) by A7, A6, AMI_3:10, SCMPDS_2:47 ; A9: (Exec ((AddTo (GBP,3,1)),s)) . (intpos 1) = s . (intpos 1) by A4, A2, A3, AMI_3:10, SCMPDS_2:48; A10: (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos 1) = (Exec (((GBP,4) := (GBP,3)),(Exec ((AddTo (GBP,3,1)),s)))) . (intpos 1) by SCMPDS_5:42 .= s . (intpos 1) by A9, A6, AMI_3:10, SCMPDS_2:47 ; A11: (Exec ((AddTo (GBP,3,1)),s)) . (intpos 3) = (s . (intpos 3)) + 1 by A4, A1, SCMPDS_2:48; A12: (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . GBP = (Exec (((GBP,4) := (GBP,3)),(Exec ((AddTo (GBP,3,1)),s)))) . GBP by SCMPDS_5:42 .= 0 by A5, A6, AMI_3:10, SCMPDS_2:47 ; then A13: DataLoc (((IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . GBP),1) = intpos (0 + 1) by SCMP_GCD:1; A14: DataLoc (((Exec ((AddTo (GBP,3,1)),s)) . GBP),3) = intpos (0 + 3) by A5, SCMP_GCD:1; A15: (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos 4) = (Exec (((GBP,4) := (GBP,3)),(Exec ((AddTo (GBP,3,1)),s)))) . (intpos 4) by SCMPDS_5:42 .= (s . (intpos 3)) + 1 by A11, A6, A14, SCMPDS_2:47 ; A16: (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos 4) = (Exec ((AddTo (GBP,1,1)),(IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)))) . (intpos 4) by SCMPDS_5:41 .= (s . (intpos 3)) + 1 by A15, A13, AMI_3:10, SCMPDS_2:48 ; A17: (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos 2) = (Exec ((AddTo (GBP,1,1)),(IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= s . (intpos 2) by A8, A13, AMI_3:10, SCMPDS_2:48 ; A18: (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos 1) = (Exec ((AddTo (GBP,1,1)),(IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= (s . (intpos 1)) + 1 by A10, A13, SCMPDS_2:48 ; A19: (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos 3) = (Exec (((GBP,4) := (GBP,3)),(Exec ((AddTo (GBP,3,1)),s)))) . (intpos 3) by SCMPDS_5:42 .= (s . (intpos 3)) + 1 by A11, A6, AMI_3:10, SCMPDS_2:47 ; A20: (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos 3) = (Exec ((AddTo (GBP,1,1)),(IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= (s . (intpos 3)) + 1 by A19, A13, AMI_3:10, SCMPDS_2:48 ; A21: (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . GBP = (Exec ((AddTo (GBP,1,1)),(IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A12, A13, AMI_3:10, SCMPDS_2:48 ; then A22: DataLoc (((IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . GBP),6) = intpos (0 + 6) by SCMP_GCD:1; thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A21, A22, AMI_3:10, SCMPDS_2:47 ; ::_thesis: ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= (s . (intpos 1)) + 1 by A18, A22, AMI_3:10, SCMPDS_2:47 ; ::_thesis: ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= s . (intpos 2) by A17, A22, AMI_3:10, SCMPDS_2:47 ; ::_thesis: ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) A23: DataLoc (((IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . GBP),1) = intpos (0 + 1) by A21, SCMP_GCD:1; thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= (s . (intpos 3)) + 1 by A20, A22, AMI_3:10, SCMPDS_2:47 ; ::_thesis: ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (s . (intpos 3)) + 1 & (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . (intpos 4) by SCMPDS_5:41 .= (s . (intpos 3)) + 1 by A16, A22, AMI_3:10, SCMPDS_2:47 ; ::_thesis: ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) ) thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . (intpos 6) by SCMPDS_5:41 .= (s . (intpos 1)) + 1 by A18, A22, A23, SCMPDS_2:47 ; ::_thesis: for i being Element of NAT st i >= 7 holds (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) A24: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_holds_ (Exec_((AddTo_(GBP,3,1)),s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 implies (Exec ((AddTo (GBP,3,1)),s)) . (intpos i) = s . (intpos i) ) assume i >= 7 ; ::_thesis: (Exec ((AddTo (GBP,3,1)),s)) . (intpos i) = s . (intpos i) then i > 3 by XXREAL_0:2; hence (Exec ((AddTo (GBP,3,1)),s)) . (intpos i) = (Initialize s) . (intpos i) by A4, A1, AMI_3:10, SCMPDS_2:48 .= s . (intpos i) by SCMPDS_5:15 ; ::_thesis: verum end; A25: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_holds_ (IExec_(((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3))),P,s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 implies (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos i) = s . (intpos i) ) assume A26: i >= 7 ; ::_thesis: (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos i) = s . (intpos i) then A27: i > 4 by XXREAL_0:2; thus (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos i) = (Exec (((GBP,4) := (GBP,3)),(Exec ((AddTo (GBP,3,1)),s)))) . (intpos i) by SCMPDS_5:42 .= (Exec ((AddTo (GBP,3,1)),s)) . (intpos i) by A6, A27, AMI_3:10, SCMPDS_2:47 .= s . (intpos i) by A24, A26 ; ::_thesis: verum end; A28: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>=_7_holds_ (IExec_((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1))),P,s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i >= 7 implies (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos i) = s . (intpos i) ) assume A29: i >= 7 ; ::_thesis: (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos i) = s . (intpos i) then A30: i > 1 by XXREAL_0:2; thus (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos i) = (Exec ((AddTo (GBP,1,1)),(IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)))) . (intpos i) by SCMPDS_5:41 .= (IExec (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))),P,s)) . (intpos i) by A13, A30, AMI_3:10, SCMPDS_2:48 .= s . (intpos i) by A25, A29 ; ::_thesis: verum end; hereby ::_thesis: verum let i be Element of NAT ; ::_thesis: ( i >= 7 implies (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) ) assume A31: i >= 7 ; ::_thesis: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = s . (intpos i) then A32: i > 6 by XXREAL_0:2; thus (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos i) = (Exec (((GBP,6) := (GBP,1)),(IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)))) . (intpos i) by SCMPDS_5:41 .= (IExec ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))),P,s)) . (intpos i) by A22, A32, AMI_3:10, SCMPDS_2:47 .= s . (intpos i) by A28, A31 ; ::_thesis: verum end; end; set jf = AddTo (GBP,2,(- 1)); set B3 = (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1))); Lm11: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 holds ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . GBP = 0 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 2)) - 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 holds ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . GBP = 0 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 2)) - 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 implies ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . GBP = 0 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 2)) - 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) ) set s1 = IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s); set Bt = IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s); set P1 = P; A1: Initialize s = s by MEMSTR_0:44; A2: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))) . GBP by SCMPDS_5:15; A3: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))) . (intpos 4) by SCMPDS_5:15; A4: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))) . (intpos 6) by SCMPDS_5:15; A5: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))) . (intpos 3) by SCMPDS_5:15; A6: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))) . (intpos 1) by SCMPDS_5:15; A7: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))) . (intpos 2) by SCMPDS_5:15; assume that A8: s . (intpos 3) >= (s . (intpos 1)) + 7 and A9: s . GBP = 0 ; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . GBP = 0 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 2)) - 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) A10: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 by A9, Lm10; A11: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 by A9, Lm10; A12: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 2) = s . (intpos 2) by A9, Lm10; ( (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6) = (s . (intpos 1)) + 1 & (s . (intpos 3)) + 1 >= (7 + (s . (intpos 1))) + 1 ) by A8, A9, Lm10, XREAL_1:6; then A13: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 4) >= 7 + ((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . (intpos 6)) by A9, Lm10; A14: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP = 0 by A9, Lm10; then A15: DataLoc (((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)) . GBP),6) = intpos (0 + 6) by SCMP_GCD:1; then ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)),P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s)),P ) by A14, A13, Lm8, A2, A3, A4; then A16: ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s),P & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s),P ) by SCMPDS_6:125, SCMPDS_6:126; A17: ( (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)) is_closed_on s,P & (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)) is_halting_on s,P ) by SCMPDS_6:20, SCMPDS_6:21; then A18: (IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . GBP = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . GBP by A16, SCMPDS_7:30 .= 0 by A14, A15, A13, Lm9, A2, A3, A4 ; then A19: DataLoc (((IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . GBP),2) = intpos (0 + 2) by SCMP_GCD:1; A20: ( ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on s,P & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on s,P ) by A16, A17, A1, SCMPDS_7:24; hence (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . GBP = (Exec ((AddTo (GBP,2,(- 1))),(IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)))) . GBP by SCMPDS_7:31 .= 0 by A18, A19, AMI_3:10, SCMPDS_2:48 ; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 2)) - 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 2) = (Exec ((AddTo (GBP,2,(- 1))),(IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)))) . (intpos 2) by A20, SCMPDS_7:31 .= ((IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . (intpos 2)) + (- 1) by A19, SCMPDS_2:48 .= ((IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . (intpos 2)) - 1 .= ((IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos 2)) - 1 by A16, A17, SCMPDS_7:30 .= (s . (intpos 2)) - 1 by A14, A12, A15, A13, Lm9, A2, A3, A4, A7 ; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) + 1 & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 3) = (Exec ((AddTo (GBP,2,(- 1))),(IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)))) . (intpos 3) by A20, SCMPDS_7:31 .= (IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . (intpos 3) by A19, AMI_3:10, SCMPDS_2:48 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos 3) by A16, A17, SCMPDS_7:30 .= (s . (intpos 3)) + 1 by A14, A11, A15, A13, Lm9, A2, A3, A4, A5 ; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (s . (intpos 1)) + 1 & ( for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) ) thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos 1) = (Exec ((AddTo (GBP,2,(- 1))),(IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)))) . (intpos 1) by A20, SCMPDS_7:31 .= (IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . (intpos 1) by A19, AMI_3:10, SCMPDS_2:48 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos 1) by A16, A17, SCMPDS_7:30 .= (s . (intpos 1)) + 1 by A14, A10, A15, A13, Lm9, A2, A3, A4, A6 ; ::_thesis: for i being Element of NAT st i <> 2 holds (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) hereby ::_thesis: verum let i be Element of NAT ; ::_thesis: ( i <> 2 implies (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) ) assume A21: i <> 2 ; ::_thesis: (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)) . (intpos i) = (Exec ((AddTo (GBP,2,(- 1))),(IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)))) . (intpos i) by A20, SCMPDS_7:31 .= (IExec ((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))),P,s)) . (intpos i) by A19, A21, AMI_3:10, SCMPDS_2:48 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),P,s))))) . (intpos i) by A16, A17, SCMPDS_7:30 ; ::_thesis: verum end; end; Lm12: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 holds ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on s,P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 holds ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on s,P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 implies ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on s,P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on s,P ) ) set a = GBP ; set b = DataLoc ((s . GBP),2); assume that A1: s . (intpos 3) >= (s . (intpos 1)) + 7 and A2: s . GBP = 0 ; ::_thesis: ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on s,P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on s,P ) A3: DataLoc ((s . GBP),2) = intpos (0 + 2) by A2, SCMP_GCD:1; now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_t_._(intpos_3)_>=_(t_._(intpos_1))_+_7_&_t_._GBP_=_s_._GBP_&_t_._(DataLoc_((s_._GBP),2))_>_0_holds_ (_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._GBP_=_t_._GBP_&_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(DataLoc_((s_._GBP),2))_=_(t_._(DataLoc_((s_._GBP),2)))_-_1_&_((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))_is_closed_on_t,Q_&_((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))_is_halting_on_t,Q_&_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(intpos_3)_>=_((IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(intpos_1))_+_7_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (intpos 3) >= (t . (intpos 1)) + 7 & t . GBP = s . GBP & t . (DataLoc ((s . GBP),2)) > 0 holds ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (intpos 3) >= (t . (intpos 1)) + 7 & t . GBP = s . GBP & t . (DataLoc ((s . GBP),2)) > 0 implies ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) ) A4: Initialize t = t by MEMSTR_0:44; assume that A5: t . (intpos 3) >= (t . (intpos 1)) + 7 and A6: t . GBP = s . GBP and t . (DataLoc ((s . GBP),2)) > 0 ; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) set t1 = IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t); set Q1 = Q; A7: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . GBP by SCMPDS_5:15; A8: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos 4) by SCMPDS_5:15; A9: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos 6) by SCMPDS_5:15; A10: (t . (intpos 3)) + 1 >= (7 + (t . (intpos 1))) + 1 by A5, XREAL_1:6; thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP by A2, A5, A6, Lm11; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 by A2, A3, A5, A6, Lm11; ::_thesis: ( ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) A11: ( (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)) is_closed_on t,Q & (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)) is_halting_on t,Q ) by SCMPDS_6:20, SCMPDS_6:21; (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6) = (t . (intpos 1)) + 1 by A2, A6, Lm10; then A12: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) >= 7 + ((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6)) by A2, A6, A10, Lm10; A13: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP = 0 by A2, A6, Lm10; then DataLoc (((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP),6) = intpos (0 + 6) by SCMP_GCD:1; then ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)),Q & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)),Q ) by A13, A12, Lm8, A8, A9, A7; then ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t),Q & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t),Q ) by SCMPDS_6:125, SCMPDS_6:126; hence ( ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q ) by A11, A4, SCMPDS_7:24; ::_thesis: (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1) = (t . (intpos 1)) + 1 by A2, A5, A6, Lm11; hence (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 by A2, A5, A6, A10, Lm11; ::_thesis: verum end; hence ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on s,P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on s,P ) by A1, Th11; ::_thesis: verum end; Lm13: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 & s . (intpos 2) > 0 holds IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s) = IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)))) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 & s . (intpos 2) > 0 holds IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s) = IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)))) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 3) >= (s . (intpos 1)) + 7 & s . GBP = 0 & s . (intpos 2) > 0 implies IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s) = IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)))) ) set a = GBP ; set b = DataLoc ((s . GBP),2); assume that A1: s . (intpos 3) >= (s . (intpos 1)) + 7 and A2: s . GBP = 0 and A3: s . (intpos 2) > 0 ; ::_thesis: IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s) = IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)))) A4: DataLoc ((s . GBP),2) = intpos (0 + 2) by A2, SCMP_GCD:1; now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_t_._(intpos_3)_>=_(t_._(intpos_1))_+_7_&_t_._GBP_=_s_._GBP_&_t_._(DataLoc_((s_._GBP),2))_>_0_holds_ (_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._GBP_=_t_._GBP_&_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(DataLoc_((s_._GBP),2))_=_(t_._(DataLoc_((s_._GBP),2)))_-_1_&_((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))_is_closed_on_t,Q_&_((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))_is_halting_on_t,Q_&_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(intpos_3)_>=_((IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(intpos_1))_+_7_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (intpos 3) >= (t . (intpos 1)) + 7 & t . GBP = s . GBP & t . (DataLoc ((s . GBP),2)) > 0 holds ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (intpos 3) >= (t . (intpos 1)) + 7 & t . GBP = s . GBP & t . (DataLoc ((s . GBP),2)) > 0 implies ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) ) A5: Initialize t = t by MEMSTR_0:44; assume that A6: t . (intpos 3) >= (t . (intpos 1)) + 7 and A7: t . GBP = s . GBP and t . (DataLoc ((s . GBP),2)) > 0 ; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) set t1 = IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t); set Q1 = Q; A8: (t . (intpos 3)) + 1 >= (7 + (t . (intpos 1))) + 1 by A6, XREAL_1:6; thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = t . GBP by A2, A6, A7, Lm11; ::_thesis: ( (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) thus (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (DataLoc ((s . GBP),2)) = (t . (DataLoc ((s . GBP),2))) - 1 by A2, A4, A6, A7, Lm11; ::_thesis: ( ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q & (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 ) A9: ( (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)) is_closed_on t,Q & (((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1)) is_halting_on t,Q ) by SCMPDS_6:20, SCMPDS_6:21; (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6) = (t . (intpos 1)) + 1 by A2, A7, Lm10; then A10: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) >= 7 + ((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6)) by A2, A7, A8, Lm10; A11: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . GBP by SCMPDS_5:15; A12: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos 4) by SCMPDS_5:15; A13: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos 6) by SCMPDS_5:15; A14: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP = 0 by A2, A7, Lm10; then DataLoc (((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP),6) = intpos (0 + 6) by SCMP_GCD:1; then ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)),Q & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)),Q ) by A14, A10, Lm8, A11, A12, A13; then ( while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_closed_on IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t),Q & while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))) is_halting_on IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t),Q ) by SCMPDS_6:125, SCMPDS_6:126; hence ( ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_closed_on t,Q & ((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))) is_halting_on t,Q ) by A9, A5, SCMPDS_7:24; ::_thesis: (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1) = (t . (intpos 1)) + 1 by A2, A6, A7, Lm11; hence (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) >= ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) + 7 by A2, A6, A7, A8, Lm11; ::_thesis: verum end; hence IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s) = IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),P,s)))) by A1, A3, A4, Th12; ::_thesis: verum end; begin theorem :: SCPISORT:14 for n, p0 being Element of NAT holds card (insert-sort (n,p0)) = 23 proof let n, p0 be Element of NAT ; ::_thesis: card (insert-sort (n,p0)) = 23 set i1 = GBP := 0; set i2 = (GBP,1) := 0; set i3 = (GBP,2) := (n - 1); set i4 = (GBP,3) := p0; thus card (insert-sort (n,p0)) = (card ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0))) + (card (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))))) by AFINSQ_1:17 .= ((card (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1)))) + 1) + (card (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))))) by SCMP_GCD:4 .= (((card ((GBP := 0) ';' ((GBP,1) := 0))) + 1) + 1) + (card (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))))) by SCMP_GCD:4 .= ((2 + 1) + 1) + (card (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))))) by SCMP_GCD:5 .= 4 + ((card (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) + 3) by SCMPDS_7:41 .= 23 by Lm5 ; ::_thesis: verum end; theorem :: SCPISORT:15 for p0, n being Element of NAT st p0 >= 7 holds insert-sort (n,p0) is parahalting proof let p0, n be Element of NAT ; ::_thesis: ( p0 >= 7 implies insert-sort (n,p0) is parahalting ) set a = GBP ; set i1 = GBP := 0; set i2 = (GBP,1) := 0; set i3 = (GBP,2) := (n - 1); set i4 = (GBP,3) := p0; set I = (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0); assume A1: p0 >= 7 ; ::_thesis: insert-sort (n,p0) is parahalting now__::_thesis:_for_s_being_State_of_SCMPDS for_P_being_Instruction-Sequence_of_SCMPDS_holds_insert-sort_(n,p0)_is_halting_on_s,P let s be State of SCMPDS; ::_thesis: for P being Instruction-Sequence of SCMPDS holds insert-sort (n,p0) is_halting_on s,P let P be Instruction-Sequence of SCMPDS; ::_thesis: insert-sort (n,p0) is_halting_on s,P set s1 = IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s)); set s2 = IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s)); set P1 = P; set s3 = IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,(Initialize s)); set s4 = Exec ((GBP := 0),(Initialize s)); A2: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . GBP = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s)))) . GBP by SCMPDS_5:15; A3: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . (intpos 1) = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s)))) . (intpos 1) by SCMPDS_5:15; A4: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . (intpos 3) = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s)))) . (intpos 3) by SCMPDS_5:15; A5: ( (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0) is_closed_on s,P & (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0) is_halting_on s,P ) by SCMPDS_6:20, SCMPDS_6:21; A6: (Exec ((GBP := 0),(Initialize s))) . GBP = 0 by SCMPDS_2:45; then A7: DataLoc (((Exec ((GBP := 0),(Initialize s))) . GBP),1) = intpos (0 + 1) by SCMP_GCD:1; A8: (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,(Initialize s))) . GBP = (Exec (((GBP,1) := 0),(Exec ((GBP := 0),(Initialize s))))) . GBP by SCMPDS_5:42 .= 0 by A6, A7, AMI_3:10, SCMPDS_2:46 ; then A9: DataLoc (((IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,(Initialize s))) . GBP),2) = intpos (0 + 2) by SCMP_GCD:1; A10: (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,(Initialize s))) . (intpos 1) = (Exec (((GBP,1) := 0),(Exec ((GBP := 0),(Initialize s))))) . (intpos 1) by SCMPDS_5:42 .= 0 by A7, SCMPDS_2:46 ; A11: (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s))) . (intpos 1) = (Exec (((GBP,2) := (n - 1)),(IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,(Initialize s))))) . (intpos 1) by SCMPDS_5:41 .= 0 by A10, A9, AMI_3:10, SCMPDS_2:46 ; A12: (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s))) . GBP = (Exec (((GBP,2) := (n - 1)),(IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,(Initialize s))))) . GBP by SCMPDS_5:41 .= 0 by A8, A9, AMI_3:10, SCMPDS_2:46 ; then A13: DataLoc (((IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s))) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; A14: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . (intpos 3) = (Exec (((GBP,3) := p0),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s))))) . (intpos 3) by SCMPDS_5:41 .= p0 by A13, SCMPDS_2:46 ; (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . (intpos 1) = (Exec (((GBP,3) := p0),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s))))) . (intpos 1) by SCMPDS_5:41 .= 0 by A11, A13, AMI_3:10, SCMPDS_2:46 ; then A15: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . (intpos 3) >= ((IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . (intpos 1)) + 7 by A1, A14; (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))) . GBP = (Exec (((GBP,3) := p0),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,(Initialize s))))) . GBP by SCMPDS_5:41 .= 0 by A12, A13, AMI_3:10, SCMPDS_2:46 ; then ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))),P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s))),P ) by A15, Lm12, A2, A3, A4; then ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s)),P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := p0)),P,(Initialize s)),P ) by SCMPDS_6:125, SCMPDS_6:126; hence insert-sort (n,p0) is_halting_on s,P by A5, SCMPDS_7:24; ::_thesis: verum end; hence insert-sort (n,p0) is parahalting by SCMPDS_6:21; ::_thesis: verum end; Lm14: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 holds IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s) = IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)))) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 holds IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s) = IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)))) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & s . (intpos 6) > 0 implies IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s) = IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)))) ) set a = GBP ; set b = DataLoc ((s . GBP),6); assume that A1: s . (intpos 4) >= 7 + (s . (intpos 6)) and A2: s . GBP = 0 and A3: s . (intpos 6) > 0 ; ::_thesis: IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s) = IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)))) A4: DataLoc ((s . GBP),6) = intpos (0 + 6) by A2, SCMP_GCD:1; A5: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_(_for_x_being_Int_position_st_x_in_{(intpos_4)}_holds_ t_._x_>=_7_+_(t_._(DataLoc_((s_._GBP),6)))_)_&_t_._GBP_=_s_._GBP_&_t_._(DataLoc_((s_._GBP),6))_>_0_holds_ (_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._GBP_=_t_._GBP_&_(((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))_is_closed_on_t,Q_&_(((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))_is_halting_on_t,Q_&_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._(DataLoc_((s_._GBP),6))_<_t_._(DataLoc_((s_._GBP),6))_&_(_for_x_being_Int_position_st_x_in_{(intpos_4)}_holds_ (IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._x_>=_7_+_((IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t))_._(DataLoc_((s_._GBP),6)))_)_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((s . GBP),6))) ) & t . GBP = s . GBP & t . (DataLoc ((s . GBP),6)) > 0 holds ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( ( for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((s . GBP),6))) ) & t . GBP = s . GBP & t . (DataLoc ((s . GBP),6)) > 0 implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) ) A6: Initialize t = t by MEMSTR_0:44; assume that A7: for x being Int_position st x in {(intpos 4)} holds t . x >= 7 + (t . (DataLoc ((s . GBP),6))) and A8: ( t . GBP = s . GBP & t . (DataLoc ((s . GBP),6)) > 0 ) ; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) set Bt = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,(Initialize t)); A9: intpos 4 in {(intpos 4)} by TARSKI:def_1; then A10: t . (intpos 4) >= 7 + (t . (intpos 6)) by A4, A7; hence (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = t . GBP by A2, A4, A8, Lm6; ::_thesis: ( (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) thus ( (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_closed_on t,Q & (((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))) is_halting_on t,Q ) by SCMPDS_6:20, SCMPDS_6:21; ::_thesis: ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) & ( for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) ) ) thus (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6)) < t . (DataLoc ((s . GBP),6)) by A2, A4, A8, A10, Lm7; ::_thesis: for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) t . (intpos 4) >= 7 + (t . (DataLoc ((s . GBP),6))) by A7, A9; then (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,(Initialize t))) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,(Initialize t))) . (intpos 6)) by A2, A4, A8, Lm7, A6; hence for x being Int_position st x in {(intpos 4)} holds (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . x >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((s . GBP),6))) by A4, A6, TARSKI:def_1; ::_thesis: verum end; for x being Int_position st x in {(intpos 4)} holds s . x >= 7 + (s . (DataLoc ((s . GBP),6))) by A1, A4, TARSKI:def_1; hence IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s) = IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),P,s)))) by A3, A4, A5, SCMPDS_8:29; ::_thesis: verum end; theorem Th16: :: SCPISORT:16 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for f, g being FinSequence of INT for k0, k being Element of NAT st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & k = s . (intpos 6) & k0 = ((s . (intpos 4)) - (s . (intpos 6))) - 1 & f is_FinSequence_on s,k0 & g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),k0 & len f = len g & len f > k & f is_non_decreasing_on 1,k holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,k + 1 & ( for i being Element of NAT st i > k + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= k + 1 holds ex j being Element of NAT st ( 1 <= j & j <= k + 1 & g . i = f . j ) ) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for f, g being FinSequence of INT for k0, k being Element of NAT st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & k = s . (intpos 6) & k0 = ((s . (intpos 4)) - (s . (intpos 6))) - 1 & f is_FinSequence_on s,k0 & g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),k0 & len f = len g & len f > k & f is_non_decreasing_on 1,k holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,k + 1 & ( for i being Element of NAT st i > k + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= k + 1 holds ex j being Element of NAT st ( 1 <= j & j <= k + 1 & g . i = f . j ) ) ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: for f, g being FinSequence of INT for k0, k being Element of NAT st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & k = s . (intpos 6) & k0 = ((s . (intpos 4)) - (s . (intpos 6))) - 1 & f is_FinSequence_on s,k0 & g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),k0 & len f = len g & len f > k & f is_non_decreasing_on 1,k holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,k + 1 & ( for i being Element of NAT st i > k + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= k + 1 holds ex j being Element of NAT st ( 1 <= j & j <= k + 1 & g . i = f . j ) ) ) let f, g be FinSequence of INT ; ::_thesis: for k0, k being Element of NAT st s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & k = s . (intpos 6) & k0 = ((s . (intpos 4)) - (s . (intpos 6))) - 1 & f is_FinSequence_on s,k0 & g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),k0 & len f = len g & len f > k & f is_non_decreasing_on 1,k holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,k + 1 & ( for i being Element of NAT st i > k + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= k + 1 holds ex j being Element of NAT st ( 1 <= j & j <= k + 1 & g . i = f . j ) ) ) let m, n be Element of NAT ; ::_thesis: ( s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & n = s . (intpos 6) & m = ((s . (intpos 4)) - (s . (intpos 6))) - 1 & f is_FinSequence_on s,m & g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),m & len f = len g & len f > n & f is_non_decreasing_on 1,n implies ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n + 1 & ( for i being Element of NAT st i > n + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= n + 1 holds ex j being Element of NAT st ( 1 <= j & j <= n + 1 & g . i = f . j ) ) ) ) assume A1: ( s . (intpos 4) >= 7 + (s . (intpos 6)) & s . GBP = 0 & n = s . (intpos 6) & m = ((s . (intpos 4)) - (s . (intpos 6))) - 1 ) ; ::_thesis: ( not f is_FinSequence_on s,m or not g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),m or not len f = len g or not len f > n or not f is_non_decreasing_on 1,n or ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n + 1 & ( for i being Element of NAT st i > n + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= n + 1 holds ex j being Element of NAT st ( 1 <= j & j <= n + 1 & g . i = f . j ) ) ) ) defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for f1, f2 being FinSequence of INT st t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & $1 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > $1 & f1 is_non_decreasing_on 1,$1 holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,$1 + 1 & ( for i being Element of NAT st i > $1 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= $1 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= $1 + 1 & f2 . i = f1 . j ) ) ); assume A2: ( f is_FinSequence_on s,m & g is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),P,s),m ) ; ::_thesis: ( not len f = len g or not len f > n or not f is_non_decreasing_on 1,n or ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n + 1 & ( for i being Element of NAT st i > n + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= n + 1 holds ex j being Element of NAT st ( 1 <= j & j <= n + 1 & g . i = f . j ) ) ) ) A3: now__::_thesis:_for_k_being_Element_of_NAT_st_S1[k]_holds_ S1[k_+_1] let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A4: S1[k] ; ::_thesis: S1[k + 1] thus S1[k + 1] ::_thesis: verum proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS for f1, f2 being FinSequence of INT st t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & k + 1 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > k + 1 & f1 is_non_decreasing_on 1,k + 1 holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: for f1, f2 being FinSequence of INT st t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & k + 1 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > k + 1 & f1 is_non_decreasing_on 1,k + 1 holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) let f1, f2 be FinSequence of INT ; ::_thesis: ( t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & k + 1 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > k + 1 & f1 is_non_decreasing_on 1,k + 1 implies ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) ) A5: Initialize t = t by MEMSTR_0:44; assume that A6: t . (intpos 4) >= 7 + (t . (intpos 6)) and A7: t . GBP = 0 and A8: k + 1 = t . (intpos 6) and A9: m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 and A10: f1 is_FinSequence_on t,m and A11: f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m and A12: len f1 = len f2 and A13: len f1 > k + 1 and A14: f1 is_non_decreasing_on 1,k + 1 ; ::_thesis: ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) set Bt = IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t); set x = DataLoc ((t . (intpos 4)),(- 1)); set y = DataLoc ((t . (intpos 4)),0); A15: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP = 0 by A6, A7, A8, Lm6; (m + 1) + (k + 1) >= 7 + (t . (intpos 6)) by A6, A8, A9; then A16: m + 1 >= 7 by A8, XREAL_1:6; A17: DataLoc ((t . (intpos 4)),(- 1)) = DataLoc (m,(k + 1)) by A8, A9 .= intpos (m + (k + 1)) by SCMP_GCD:1 ; A18: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 6) by SCMPDS_5:15; A19: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4) = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos 4) by SCMPDS_5:15; A20: (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . GBP = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP by SCMPDS_5:15; A21: ( t . (DataLoc ((t . (intpos 4)),(- 1))) > t . (DataLoc ((t . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((t . (intpos 4)),(- 1))) = t . (DataLoc ((t . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((t . (intpos 4)),0)) = t . (DataLoc ((t . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6) = (t . (intpos 6)) - 1 & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4) = (t . (intpos 4)) - 1 ) ) by A6, A7, A8, Lm7; A22: ( t . (DataLoc ((t . (intpos 4)),(- 1))) <= t . (DataLoc ((t . (intpos 4)),0)) implies ( (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((t . (intpos 4)),(- 1))) = t . (DataLoc ((t . (intpos 4)),(- 1))) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc ((t . (intpos 4)),0)) = t . (DataLoc ((t . (intpos 4)),0)) & (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6) = 0 ) ) by A6, A7, A8, Lm7; A23: DataLoc ((t . (intpos 4)),0) = intpos (m + (k + 2)) by A8, A9, SCMP_GCD:1; A24: (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4) >= 7 + ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6)) by A6, A7, A8, Lm7; percases ( t . (DataLoc ((t . (intpos 4)),(- 1))) > t . (DataLoc ((t . (intpos 4)),0)) or t . (DataLoc ((t . (intpos 4)),(- 1))) <= t . (DataLoc ((t . (intpos 4)),0)) ) ; supposeA25: t . (DataLoc ((t . (intpos 4)),(- 1))) > t . (DataLoc ((t . (intpos 4)),0)) ; ::_thesis: ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_f2_holds_ f2_._i_=_(IExec_((while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))),Q,(Initialize_(IExec_(((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))),Q,t)))))_._(intpos_(m_+_i)) let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len f2 implies f2 . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos (m + i)) ) assume ( 1 <= i & i <= len f2 ) ; ::_thesis: f2 . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos (m + i)) hence f2 . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t)) . (intpos (m + i)) by A11, Def1 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos (m + i)) by A6, A7, A8, Lm14 ; ::_thesis: verum end; then A26: f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)))),m by Def1; A27: k + 1 < k + 2 by XREAL_1:6; consider h being FinSequence of INT such that A28: len h = len f1 and A29: for i being Element of NAT st 1 <= i & i <= len h holds h . i = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos (m + i)) by Th1; k + 1 > k by XREAL_1:29; then A30: len h > k by A13, A28, XXREAL_0:2; A31: now__::_thesis:_for_i_being_Element_of_NAT_st_i_<>_k_+_1_&_i_<>_k_+_2_&_1_<=_i_&_i_<=_len_f1_holds_ h_._i_=_f1_._i let i be Element of NAT ; ::_thesis: ( i <> k + 1 & i <> k + 2 & 1 <= i & i <= len f1 implies h . i = f1 . i ) assume that A32: ( i <> k + 1 & i <> k + 2 ) and A33: 1 <= i and A34: i <= len f1 ; ::_thesis: h . i = f1 . i A35: ( m + i <> (t . (intpos 4)) - 1 & m + i <> t . (intpos 4) ) by A8, A9, A32; m + i >= m + 1 by A33, XREAL_1:6; then A36: m + i >= 7 by A16, XXREAL_0:2; thus h . i = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos (m + i)) by A28, A29, A33, A34 .= t . (intpos (m + i)) by A6, A7, A8, A36, A35, Lm7 .= f1 . i by A10, A33, A34, Def1 ; ::_thesis: verum end; now__::_thesis:_for_i,_j_being_Element_of_NAT_st_1_<=_i_&_i_<=_j_&_j_<=_k_holds_ h_._i_<=_h_._j let i, j be Element of NAT ; ::_thesis: ( 1 <= i & i <= j & j <= k implies h . i <= h . j ) assume that A37: 1 <= i and A38: i <= j and A39: j <= k ; ::_thesis: h . i <= h . j A40: j <= len f1 by A28, A30, A39, XXREAL_0:2; then A41: i <= len f1 by A38, XXREAL_0:2; A42: k < k + 1 by XREAL_1:29; then A43: j < k + 1 by A39, XXREAL_0:2; k + 1 < (k + 1) + 1 by XREAL_1:29; then A44: k < (k + 1) + 1 by A42, XXREAL_0:2; j >= 1 by A37, A38, XXREAL_0:2; then A45: h . j = f1 . j by A31, A39, A42, A44, A40; j < k + 2 by A39, A44, XXREAL_0:2; then h . i = f1 . i by A31, A37, A38, A43, A41; hence h . i <= h . j by A14, A37, A38, A43, A45, GRAPH_2:def_12; ::_thesis: verum end; then A46: h is_non_decreasing_on 1,k by GRAPH_2:def_12; A47: len f1 >= (k + 1) + 1 by A13, INT_1:7; A48: 1 <= k + 1 by NAT_1:11; then A49: 1 <= k + 2 by A27, XXREAL_0:2; then A50: h . (k + 2) = t . (DataLoc ((t . (intpos 4)),(- 1))) by A23, A21, A25, A28, A29, A47; then A51: h . (k + 2) = f1 . (k + 1) by A10, A13, A17, A48, Def1; A52: (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 4)) - ((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos 6))) - 1 = m by A9, A21, A25; A53: h . (k + 1) = t . (DataLoc ((t . (intpos 4)),0)) by A13, A17, A21, A25, A28, A29, NAT_1:11; then h . (k + 1) = f1 . (k + 2) by A10, A23, A49, A47, Def1; then A54: f1,h are_fiberwise_equipotent by A13, A28, A31, A48, A49, A47, A51, Th3; A55: h is_FinSequence_on Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)),m proof let i be Element of NAT ; :: according to SCPISORT:def_1 ::_thesis: ( 1 <= i & i <= len h implies h . i = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos (m + i)) ) assume ( 1 <= i & i <= len h ) ; ::_thesis: h . i = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos (m + i)) then h . i = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos (m + i)) by A29; hence h . i = (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos (m + i)) by SCMPDS_5:15; ::_thesis: verum end; h,f2 are_fiberwise_equipotent by A4, A8, A12, A15, A24, A21, A25, A28, A52, A26, A30, A46, A55, A18, A19, A20; hence f1,f2 are_fiberwise_equipotent by A54, CLASSES1:76; ::_thesis: ( f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) A56: f2 is_non_decreasing_on 1,k + 1 by A4, A8, A12, A15, A24, A21, A25, A28, A52, A55, A18, A19, A20, A26, A30, A46; now__::_thesis:_for_i,_j_being_Element_of_NAT_st_1_<=_i_&_i_<=_j_&_j_<=_(k_+_1)_+_1_holds_ f2_._i_<=_f2_._j let i, j be Element of NAT ; ::_thesis: ( 1 <= i & i <= j & j <= (k + 1) + 1 implies f2 . b1 <= f2 . b2 ) assume that A57: 1 <= i and A58: i <= j and A59: j <= (k + 1) + 1 ; ::_thesis: f2 . b1 <= f2 . b2 percases ( j <= k + 1 or j = (k + 1) + 1 ) by A59, NAT_1:8; suppose j <= k + 1 ; ::_thesis: f2 . b1 <= f2 . b2 hence f2 . i <= f2 . j by A56, A57, A58, GRAPH_2:def_12; ::_thesis: verum end; supposeA60: j = (k + 1) + 1 ; ::_thesis: f2 . b1 <= f2 . b2 hereby ::_thesis: verum percases ( i = j or i <> j ) ; suppose i = j ; ::_thesis: f2 . i <= f2 . j hence f2 . i <= f2 . j ; ::_thesis: verum end; suppose i <> j ; ::_thesis: f2 . i <= f2 . j then i < j by A58, XXREAL_0:1; then i <= k + 1 by A60, NAT_1:13; then consider mm being Element of NAT such that A61: 1 <= mm and A62: mm <= k + 1 and A63: f2 . i = h . mm by A4, A8, A12, A15, A24, A21, A25, A28, A52, A55, A26, A30, A46, A57, A18, A19, A20; A64: f2 . j = h . (k + 2) by A4, A8, A12, A15, A24, A21, A25, A28, A52, A55, A26, A30, A46, A27, A47, A60, A18, A19, A20; hereby ::_thesis: verum percases ( mm = k + 1 or mm <> k + 1 ) ; suppose mm = k + 1 ; ::_thesis: f2 . i <= f2 . j hence f2 . i <= f2 . j by A13, A17, A21, A22, A28, A29, A50, A61, A63, A64; ::_thesis: verum end; supposeA65: mm <> k + 1 ; ::_thesis: f2 . i <= f2 . j mm < k + 2 by A27, A62, XXREAL_0:2; then mm < len h by A28, A47, XXREAL_0:2; then A66: h . mm = f1 . mm by A28, A31, A27, A61, A62, A65; f2 . j = f1 . (k + 1) by A10, A13, A17, A48, A50, A64, Def1; hence f2 . i <= f2 . j by A14, A61, A62, A63, A66, GRAPH_2:def_12; ::_thesis: verum end; end; end; end; end; end; end; end; end; hence f2 is_non_decreasing_on 1,(k + 1) + 1 by GRAPH_2:def_12; ::_thesis: ( ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) hereby ::_thesis: for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) let i be Element of NAT ; ::_thesis: ( i > (k + 1) + 1 & i <= len f1 implies f2 . i = f1 . i ) assume that A67: i > (k + 1) + 1 and A68: i <= len f1 ; ::_thesis: f2 . i = f1 . i A69: k + 1 < (k + 1) + 1 by XREAL_1:29; then A70: i > k + 1 by A67, XXREAL_0:2; 1 <= k + 1 by NAT_1:11; then A71: 1 <= i by A70, XXREAL_0:2; thus f2 . i = h . i by A4, A8, A12, A15, A24, A21, A25, A28, A52, A55, A26, A30, A46, A68, A70, A18, A19, A20 .= f1 . i by A31, A67, A68, A69, A71 ; ::_thesis: verum end; hereby ::_thesis: verum let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= (k + 1) + 1 implies ex j being Element of NAT st ( 1 <= b2 & b2 <= (k + 1) + 1 & f2 . j = f1 . b2 ) ) assume that A72: 1 <= i and A73: i <= (k + 1) + 1 ; ::_thesis: ex j being Element of NAT st ( 1 <= b2 & b2 <= (k + 1) + 1 & f2 . j = f1 . b2 ) percases ( i = (k + 1) + 1 or i <> (k + 1) + 1 ) ; supposeA74: i = (k + 1) + 1 ; ::_thesis: ex j being Element of NAT st ( 1 <= b2 & b2 <= (k + 1) + 1 & f2 . j = f1 . b2 ) take j = k + 1; ::_thesis: ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) thus 1 <= j by NAT_1:11; ::_thesis: ( j <= (k + 1) + 1 & f2 . i = f1 . j ) thus j <= (k + 1) + 1 by NAT_1:11; ::_thesis: f2 . i = f1 . j thus f2 . i = f1 . j by A4, A8, A12, A15, A24, A21, A25, A28, A52, A55, A26, A30, A46, A27, A47, A51, A74, A18, A19, A20; ::_thesis: verum end; suppose i <> (k + 1) + 1 ; ::_thesis: ex j being Element of NAT st ( 1 <= b2 & b2 <= (k + 1) + 1 & f2 . j = f1 . b2 ) then i < (k + 1) + 1 by A73, XXREAL_0:1; then i <= k + 1 by NAT_1:13; then consider mm being Element of NAT such that A75: 1 <= mm and A76: mm <= k + 1 and A77: f2 . i = h . mm by A4, A8, A12, A15, A24, A21, A25, A28, A52, A55, A26, A30, A46, A72, A18, A19, A20; hereby ::_thesis: verum A78: k + 2 = (k + 1) + 1 ; percases ( mm = k + 1 or mm <> k + 1 ) ; supposeA79: mm = k + 1 ; ::_thesis: ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) take j = k + 2; ::_thesis: ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) thus 1 <= j by A78, NAT_1:11; ::_thesis: ( j <= (k + 1) + 1 & f2 . i = f1 . j ) thus j <= (k + 1) + 1 ; ::_thesis: f2 . i = f1 . j thus f2 . i = f1 . j by A10, A23, A49, A47, A53, A77, A79, Def1; ::_thesis: verum end; supposeA80: mm <> k + 1 ; ::_thesis: ex mm being Element of NAT st ( 1 <= mm & mm <= (k + 1) + 1 & f2 . i = f1 . mm ) take mm = mm; ::_thesis: ( 1 <= mm & mm <= (k + 1) + 1 & f2 . i = f1 . mm ) thus 1 <= mm by A75; ::_thesis: ( mm <= (k + 1) + 1 & f2 . i = f1 . mm ) thus mm <= (k + 1) + 1 by A27, A76, XXREAL_0:2; ::_thesis: f2 . i = f1 . mm mm < k + 2 by A27, A76, XXREAL_0:2; then mm < len f1 by A47, XXREAL_0:2; hence f2 . i = f1 . mm by A31, A27, A75, A76, A77, A80; ::_thesis: verum end; end; end; end; end; end; end; supposeA81: t . (DataLoc ((t . (intpos 4)),(- 1))) <= t . (DataLoc ((t . (intpos 4)),0)) ; ::_thesis: ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) A82: now__::_thesis:_for_i_being_Nat_st_i_>=_1_&_i_<=_len_f1_holds_ f1_._i_=_f2_._i let i be Nat; ::_thesis: ( i >= 1 & i <= len f1 implies f1 . b1 = f2 . b1 ) assume that A83: i >= 1 and A84: i <= len f1 ; ::_thesis: f1 . b1 = f2 . b1 A85: i in NAT by ORDINAL1:def_12; then A86: f1 . i = t . (intpos (m + i)) by A10, A83, A84, Def1; (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP),6)) = 0 by A15, A22, A81, SCMP_GCD:1; then A87: (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (DataLoc (((IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . GBP),6)) = 0 by SCMPDS_5:15; m + i >= m + 1 by A83, XREAL_1:6; then A88: m + i >= 7 by A16, XXREAL_0:2; A89: (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (intpos (m + i)) = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos (m + i)) by SCMPDS_5:15; A90: (Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . (DataLoc (((Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))) . GBP),6)) <= 0 by A87, SCMPDS_5:15; percases ( m + i = (t . (intpos 4)) - 1 or m + i = t . (intpos 4) or ( m + i <> (t . (intpos 4)) - 1 & m + i <> t . (intpos 4) ) ) ; supposeA91: m + i = (t . (intpos 4)) - 1 ; ::_thesis: f1 . b1 = f2 . b1 hence f1 . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (DataLoc ((t . (intpos 4)),(- 1))) by A8, A9, A17, A22, A81, A86, A89, A90, SCMPDS_8:23 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize t))) . (DataLoc ((t . (intpos 4)),(- 1))) by A6, A7, A8, Lm14, A5 .= f2 . i by A8, A9, A11, A12, A13, A17, A83, A91, Def1, A5 ; ::_thesis: verum end; supposeA92: m + i = t . (intpos 4) ; ::_thesis: f1 . b1 = f2 . b1 hence f1 . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (DataLoc ((t . (intpos 4)),0)) by A8, A9, A23, A22, A81, A86, A89, A90, SCMPDS_8:23 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize t))) . (DataLoc ((t . (intpos 4)),0)) by A6, A7, A8, Lm14, A5 .= f2 . i by A8, A9, A11, A12, A23, A83, A84, A92, Def1, A5 ; ::_thesis: verum end; suppose ( m + i <> (t . (intpos 4)) - 1 & m + i <> t . (intpos 4) ) ; ::_thesis: f1 . b1 = f2 . b1 hence f1 . i = (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t)) . (intpos (m + i)) by A6, A7, A8, A88, A86, Lm7 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))),Q,t))))) . (intpos (m + i)) by A89, A90, SCMPDS_8:23 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize t))) . (intpos (m + i)) by A6, A7, A8, Lm14, A5 .= f2 . i by A11, A12, A85, A83, A84, Def1, A5 ; ::_thesis: verum end; end; end; then A93: f1 = f2 by A12, FINSEQ_1:14; thus f1,f2 are_fiberwise_equipotent by A12, A82, FINSEQ_1:14; ::_thesis: ( f2 is_non_decreasing_on 1,(k + 1) + 1 & ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) now__::_thesis:_for_i,_j_being_Element_of_NAT_st_1_<=_i_&_i_<=_j_&_j_<=_(k_+_1)_+_1_holds_ f1_._i_<=_f1_._j let i, j be Element of NAT ; ::_thesis: ( 1 <= i & i <= j & j <= (k + 1) + 1 implies f1 . b1 <= f1 . b2 ) assume that A94: 1 <= i and A95: i <= j and A96: j <= (k + 1) + 1 ; ::_thesis: f1 . b1 <= f1 . b2 percases ( j <= k + 1 or j = (k + 1) + 1 ) by A96, NAT_1:8; suppose j <= k + 1 ; ::_thesis: f1 . b1 <= f1 . b2 hence f1 . i <= f1 . j by A14, A94, A95, GRAPH_2:def_12; ::_thesis: verum end; supposeA97: j = (k + 1) + 1 ; ::_thesis: f1 . b1 <= f1 . b2 hereby ::_thesis: verum percases ( i = j or i <> j ) ; suppose i = j ; ::_thesis: f1 . i <= f1 . j hence f1 . i <= f1 . j ; ::_thesis: verum end; suppose i <> j ; ::_thesis: f1 . i <= f1 . j then i < j by A95, XXREAL_0:1; then i <= k + 1 by A97, NAT_1:13; then A98: f1 . i <= f1 . (k + 1) by A14, A94, GRAPH_2:def_12; 1 <= k + 1 by NAT_1:11; then A99: f1 . (k + 1) = t . (DataLoc ((t . (intpos 4)),(- 1))) by A10, A13, A17, Def1; ( 1 <= (k + 1) + 1 & j <= len f1 ) by A13, A97, INT_1:7, NAT_1:11; then f1 . j = t . (DataLoc ((t . (intpos 4)),0)) by A10, A23, A97, Def1; hence f1 . i <= f1 . j by A81, A98, A99, XXREAL_0:2; ::_thesis: verum end; end; end; end; end; end; hence f2 is_non_decreasing_on 1,(k + 1) + 1 by A93, GRAPH_2:def_12; ::_thesis: ( ( for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) ) ) thus for i being Element of NAT st i > (k + 1) + 1 & i <= len f1 holds f1 . i = f2 . i by A12, A82, FINSEQ_1:14; ::_thesis: for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) thus for i being Element of NAT st 1 <= i & i <= (k + 1) + 1 holds ex j being Element of NAT st ( 1 <= j & j <= (k + 1) + 1 & f2 . i = f1 . j ) by A93; ::_thesis: verum end; end; end; end; A100: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS for f1, f2 being FinSequence of INT st t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & 0 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > 0 & f1 is_non_decreasing_on 1, 0 holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,0 + 1 & ( for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: for f1, f2 being FinSequence of INT st t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & 0 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > 0 & f1 is_non_decreasing_on 1, 0 holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,0 + 1 & ( for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ) ) let f1, f2 be FinSequence of INT ; ::_thesis: ( t . (intpos 4) >= 7 + (t . (intpos 6)) & t . GBP = 0 & 0 = t . (intpos 6) & m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 & f1 is_FinSequence_on t,m & f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m & len f1 = len f2 & len f1 > 0 & f1 is_non_decreasing_on 1, 0 implies ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,0 + 1 & ( for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ) ) ) A101: Initialize t = t by MEMSTR_0:44; assume that t . (intpos 4) >= 7 + (t . (intpos 6)) and A102: ( t . GBP = 0 & 0 = t . (intpos 6) ) and m = ((t . (intpos 4)) - (t . (intpos 6))) - 1 and A103: f1 is_FinSequence_on t,m and A104: f2 is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,t),m and A105: len f1 = len f2 and len f1 > 0 and f1 is_non_decreasing_on 1, 0 ; ::_thesis: ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,0 + 1 & ( for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ) ) A106: t . (DataLoc ((t . GBP),6)) = 0 by A102, SCMP_GCD:1; A107: now__::_thesis:_for_i_being_Nat_st_1_<=_i_&_i_<=_len_f1_holds_ f1_._i_=_f2_._i let i be Nat; ::_thesis: ( 1 <= i & i <= len f1 implies f1 . i = f2 . i ) assume A108: ( 1 <= i & i <= len f1 ) ; ::_thesis: f1 . i = f2 . i A109: i in NAT by ORDINAL1:def_12; hence f1 . i = t . (intpos (m + i)) by A103, A108, Def1 .= (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize t))) . (intpos (m + i)) by A106, A101, SCMPDS_8:23 .= f2 . i by A104, A105, A109, A108, Def1, A101 ; ::_thesis: verum end; hence f1,f2 are_fiberwise_equipotent by A105, FINSEQ_1:14; ::_thesis: ( f2 is_non_decreasing_on 1,0 + 1 & ( for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ) ) thus f2 is_non_decreasing_on 1,0 + 1 by GRAPH_2:63; ::_thesis: ( ( for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i ) & ( for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ) ) thus for i being Element of NAT st i > 0 + 1 & i <= len f1 holds f1 . i = f2 . i by A107; ::_thesis: for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) f1 = f2 by A105, A107, FINSEQ_1:14; hence for i being Element of NAT st 1 <= i & i <= 0 + 1 holds ex j being Element of NAT st ( 1 <= j & j <= 0 + 1 & f2 . i = f1 . j ) ; ::_thesis: verum end; A110: for k being Element of NAT holds S1[k] from NAT_1:sch_1(A100, A3); assume ( len f = len g & len f > n & f is_non_decreasing_on 1,n ) ; ::_thesis: ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n + 1 & ( for i being Element of NAT st i > n + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= n + 1 holds ex j being Element of NAT st ( 1 <= j & j <= n + 1 & g . i = f . j ) ) ) hence ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n + 1 & ( for i being Element of NAT st i > n + 1 & i <= len f holds f . i = g . i ) & ( for i being Element of NAT st 1 <= i & i <= n + 1 holds ex j being Element of NAT st ( 1 <= j & j <= n + 1 & g . i = f . j ) ) ) by A1, A2, A110; ::_thesis: verum end; Lm15: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for f, g being FinSequence of INT for p0, n being Element of NAT st s . GBP = 0 & s . (intpos 2) = n - 1 & s . (intpos 3) = p0 + 1 & s . (intpos 1) = 0 & p0 >= 6 & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s),p0 & len f = n & len g = n holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for f, g being FinSequence of INT for p0, n being Element of NAT st s . GBP = 0 & s . (intpos 2) = n - 1 & s . (intpos 3) = p0 + 1 & s . (intpos 1) = 0 & p0 >= 6 & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s),p0 & len f = n & len g = n holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: for f, g being FinSequence of INT for p0, n being Element of NAT st s . GBP = 0 & s . (intpos 2) = n - 1 & s . (intpos 3) = p0 + 1 & s . (intpos 1) = 0 & p0 >= 6 & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s),p0 & len f = n & len g = n holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) let f, g be FinSequence of INT ; ::_thesis: for p0, n being Element of NAT st s . GBP = 0 & s . (intpos 2) = n - 1 & s . (intpos 3) = p0 + 1 & s . (intpos 1) = 0 & p0 >= 6 & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s),p0 & len f = n & len g = n holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) let p0, n be Element of NAT ; ::_thesis: ( s . GBP = 0 & s . (intpos 2) = n - 1 & s . (intpos 3) = p0 + 1 & s . (intpos 1) = 0 & p0 >= 6 & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s),p0 & len f = n & len g = n implies ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) ) assume that A1: s . GBP = 0 and A2: s . (intpos 2) = n - 1 and A3: s . (intpos 3) = p0 + 1 and A4: s . (intpos 1) = 0 and A5: p0 >= 6 and A6: ( f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,s),p0 ) and A7: len f = n and A8: len g = n ; ::_thesis: ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) percases ( n = 0 or n <> 0 ) ; supposeA9: n = 0 ; ::_thesis: ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) then g = {} by A8; hence f,g are_fiberwise_equipotent by A7, A9; ::_thesis: g is_non_decreasing_on 1,n thus g is_non_decreasing_on 1,n by A9, GRAPH_2:63; ::_thesis: verum end; suppose n <> 0 ; ::_thesis: ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) then n >= 1 + 0 by INT_1:7; then n - 1 >= 0 by XREAL_1:19; then reconsider n1 = n - 1 as Element of NAT by INT_1:3; defpred S1[ Nat] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for f1, f2 being FinSequence of INT for m being Element of NAT st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = $1 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ); A10: ( (s . (intpos 2)) + (s . (intpos 1)) = (n - 1) + 0 & 1 = n - (s . (intpos 2)) ) by A2, A4; A11: now__::_thesis:_for_k_being_Element_of_NAT_st_S1[k]_holds_ S1[k_+_1] let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A12: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_f1,_f2_being_FinSequence_of_INT_ for_m_being_Element_of_NAT_ for_Q_being_Instruction-Sequence_of_SCMPDS_st_t_._GBP_=_0_&_(t_._(intpos_2))_+_(t_._(intpos_1))_=_n_-_1_&_t_._(intpos_2)_=_k_+_1_&_m_=_n_-_(t_._(intpos_2))_&_p0_=_((t_._(intpos_3))_-_(t_._(intpos_1)))_-_1_&_f1_is_FinSequence_on_t,p0_&_f2_is_FinSequence_on_IExec_((for-down_(GBP,2,1,(((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))))),Q,t),p0_&_f1_is_non_decreasing_on_1,m_&_len_f1_=_n_&_len_f2_=_n_holds_ (_f1,f2_are_fiberwise_equipotent_&_f2_is_non_decreasing_on_1,n_) let t be 0 -started State of SCMPDS; ::_thesis: for f1, f2 being FinSequence of INT for m being Element of NAT for Q being Instruction-Sequence of SCMPDS st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = k + 1 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) let f1, f2 be FinSequence of INT ; ::_thesis: for m being Element of NAT for Q being Instruction-Sequence of SCMPDS st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = k + 1 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) let m be Element of NAT ; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = k + 1 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = k + 1 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n implies ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) ) assume that A13: t . GBP = 0 and A14: (t . (intpos 2)) + (t . (intpos 1)) = n - 1 and A15: t . (intpos 2) = k + 1 and A16: m = n - (t . (intpos 2)) and A17: p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 and A18: f1 is_FinSequence_on t,p0 and A19: f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 and A20: ( f1 is_non_decreasing_on 1,m & len f1 = n ) and A21: len f2 = n ; ::_thesis: ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) set t1 = IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t); set Bt = IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t); set Q1 = Q; A22: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . GBP by SCMPDS_5:15; A23: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos 4) by SCMPDS_5:15; A24: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6) = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos 6) by SCMPDS_5:15; A25: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) = (t . (intpos 3)) + 1 by A13, Lm10; p0 + ((t . (intpos 1)) + 1) = t . (intpos 3) by A17; then t . (intpos 3) >= 6 + ((t . (intpos 1)) + 1) by A5, XREAL_1:6; then A26: t . (intpos 3) >= (6 + 1) + (t . (intpos 1)) ; then A27: (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP = 0 by A13, Lm11; A28: (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 2) = (t . (intpos 2)) - 1 by A13, A26, Lm11; then A29: n - ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 2)) = m + 1 by A16; A30: (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1) = (t . (intpos 1)) + 1 by A13, A26, Lm11; then A31: ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 2)) + ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1)) = n - 1 by A14, A28; (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) + 1 by A13, A26, Lm11; then A32: (((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3)) - ((IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1))) - 1 = p0 by A17, A30; A33: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6) = (t . (intpos 1)) + 1 by A13, Lm10; then A34: p0 = (((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4)) - ((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6))) - 1 by A17, A25; now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_f2_holds_ f2_._i_=_(IExec_((for-down_(GBP,2,1,(((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))))),Q,(Initialize_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t)))))_._(intpos_(p0_+_i)) let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len f2 implies f2 . i = (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))))) . (intpos (p0 + i)) ) assume ( 1 <= i & i <= len f2 ) ; ::_thesis: f2 . i = (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))))) . (intpos (p0 + i)) hence f2 . i = (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t)) . (intpos (p0 + i)) by A19, Def1 .= (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))))) . (intpos (p0 + i)) by A13, A15, A26, Lm13 ; ::_thesis: verum end; then A35: f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,(Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)))),p0 by Def1; A36: now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_f1_holds_ f1_._i_=_(IExec_(((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1))),Q,t))_._(intpos_(p0_+_i)) A37: p0 + 1 >= 6 + 1 by A5, XREAL_1:6; let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len f1 implies f1 . i = (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos (p0 + i)) ) assume that A38: 1 <= i and A39: i <= len f1 ; ::_thesis: f1 . i = (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos (p0 + i)) p0 + 1 <= p0 + i by A38, XREAL_1:6; then A40: p0 + i >= 7 by A37, XXREAL_0:2; thus f1 . i = t . (intpos (p0 + i)) by A18, A38, A39, Def1 .= (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos (p0 + i)) by A13, A40, Lm10 ; ::_thesis: verum end; (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) = p0 + (((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6)) + 1) by A17, A25, A33; then (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) >= 6 + (((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6)) + 1) by A5, XREAL_1:6; then A41: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 4) >= (6 + 1) + ((IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos 6)) ; m + (k + 1) = n by A15, A16; then A42: n > 0 + m by XREAL_1:6; consider h being FinSequence of INT such that A43: len h = n and A44: for i being Element of NAT st 1 <= i & i <= len h holds h . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))))) . (intpos (p0 + i)) by Th1; A45: h is_FinSequence_on IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)))),p0 by A44, Def1; A46: now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_h_holds_ h_._i_=_(IExec_(((((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0))))))))_';'_(AddTo_(GBP,2,(-_1)))),Q,t))_._(intpos_(p0_+_i)) A47: p0 + 1 >= 6 + 1 by A5, XREAL_1:6; let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len h implies h . i = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos (p0 + i)) ) assume that A48: 1 <= i and A49: i <= len h ; ::_thesis: h . i = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos (p0 + i)) p0 + 1 <= p0 + i by A48, XREAL_1:6; then p0 + i >= 7 by A47, XXREAL_0:2; then A50: p0 + i > 2 by XXREAL_0:2; thus h . i = (IExec ((while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))),Q,(Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))))) . (intpos (p0 + i)) by A44, A48, A49 .= (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos (p0 + i)) by A13, A26, A50, Lm11 ; ::_thesis: verum end; A51: f1 is_FinSequence_on Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)),p0 proof let i be Element of NAT ; :: according to SCPISORT:def_1 ::_thesis: ( 1 <= i & i <= len f1 implies f1 . i = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos (p0 + i)) ) assume ( 1 <= i & i <= len f1 ) ; ::_thesis: f1 . i = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos (p0 + i)) then f1 . i = (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . (intpos (p0 + i)) by A36; hence f1 . i = (Initialize (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t))) . (intpos (p0 + i)) by SCMPDS_5:15; ::_thesis: verum end; A52: (IExec (((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))),Q,t)) . GBP = 0 by A13, Lm10; then A53: f1,h are_fiberwise_equipotent by A14, A16, A17, A20, A25, A43, A34, A41, A45, A42, Th16, A23, A24, A22, A51; A54: h is_non_decreasing_on 1,m + 1 by A14, A16, A17, A20, A25, A43, A34, A41, A45, A42, Th16, A23, A24, A22, A51, A52; A55: (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . GBP = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . GBP by SCMPDS_5:15; A56: (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . (intpos 1) = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 1) by SCMPDS_5:15; A57: (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . (intpos 2) = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 2) by SCMPDS_5:15; A58: (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . (intpos 3) = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos 3) by SCMPDS_5:15; A59: h is_FinSequence_on Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)),p0 proof let i be Element of NAT ; :: according to SCPISORT:def_1 ::_thesis: ( 1 <= i & i <= len h implies h . i = (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . (intpos (p0 + i)) ) assume A60: ( 1 <= i & i <= len h ) ; ::_thesis: h . i = (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . (intpos (p0 + i)) thus h . i = (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t)) . (intpos (p0 + i)) by A60, A46 .= (Initialize (IExec (((((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))) ';' (AddTo (GBP,2,(- 1)))),Q,t))) . (intpos (p0 + i)) by SCMPDS_5:15 ; ::_thesis: verum end; then h,f2 are_fiberwise_equipotent by A12, A15, A16, A21, A43, A27, A31, A29, A32, A35, A54, A55, A56, A57, A58; hence f1,f2 are_fiberwise_equipotent by A53, CLASSES1:76; ::_thesis: f2 is_non_decreasing_on 1,n thus f2 is_non_decreasing_on 1,n by A12, A15, A16, A21, A43, A54, A27, A31, A29, A32, A35, A55, A56, A57, A58, A59; ::_thesis: verum end; hence S1[k + 1] ; ::_thesis: verum end; A61: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS for f1, f2 being FinSequence of INT for m being Element of NAT st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = 0 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: for f1, f2 being FinSequence of INT for m being Element of NAT st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = 0 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) let f1, f2 be FinSequence of INT ; ::_thesis: for m being Element of NAT st t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = 0 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n holds ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) let m be Element of NAT ; ::_thesis: ( t . GBP = 0 & (t . (intpos 2)) + (t . (intpos 1)) = n - 1 & t . (intpos 2) = 0 & m = n - (t . (intpos 2)) & p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 & f1 is_FinSequence_on t,p0 & f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 & f1 is_non_decreasing_on 1,m & len f1 = n & len f2 = n implies ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) ) A62: Initialize t = t by MEMSTR_0:44; assume that A63: t . GBP = 0 and (t . (intpos 2)) + (t . (intpos 1)) = n - 1 and A64: t . (intpos 2) = 0 and A65: m = n - (t . (intpos 2)) and p0 = ((t . (intpos 3)) - (t . (intpos 1))) - 1 and A66: f1 is_FinSequence_on t,p0 and A67: f2 is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,t),p0 and A68: f1 is_non_decreasing_on 1,m and A69: ( len f1 = n & len f2 = n ) ; ::_thesis: ( f1,f2 are_fiberwise_equipotent & f2 is_non_decreasing_on 1,n ) A70: t . (DataLoc ((t . GBP),2)) = 0 by A63, A64, SCMP_GCD:1; A71: now__::_thesis:_for_i_being_Nat_st_1_<=_i_&_i_<=_len_f2_holds_ f2_._i_=_f1_._i let i be Nat; ::_thesis: ( 1 <= i & i <= len f2 implies f2 . i = f1 . i ) assume A72: ( 1 <= i & i <= len f2 ) ; ::_thesis: f2 . i = f1 . i A73: i in NAT by ORDINAL1:def_12; hence f2 . i = (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),Q,(Initialize t))) . (intpos (p0 + i)) by A67, A72, Def1, A62 .= t . (intpos (p0 + i)) by A70, SCMPDS_7:47 .= f1 . i by A66, A69, A73, A72, Def1 ; ::_thesis: verum end; hence f1,f2 are_fiberwise_equipotent by A69, FINSEQ_1:14; ::_thesis: f2 is_non_decreasing_on 1,n thus f2 is_non_decreasing_on 1,n by A64, A65, A68, A69, A71, FINSEQ_1:14; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A61, A11); then A74: S1[n1] ; ( p0 = ((s . (intpos 3)) - (s . (intpos 1))) - 1 & f is_non_decreasing_on 1,1 ) by A3, A4, GRAPH_2:63; hence ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) by A1, A6, A7, A8, A10, A74; ::_thesis: verum end; end; end; theorem :: SCPISORT:17 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for f, g being FinSequence of INT for p0, n being Element of NAT st p0 >= 6 & len f = n & len g = n & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((insert-sort (n,(p0 + 1))),P,s),p0 holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for f, g being FinSequence of INT for p0, n being Element of NAT st p0 >= 6 & len f = n & len g = n & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((insert-sort (n,(p0 + 1))),P,s),p0 holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: for f, g being FinSequence of INT for p0, n being Element of NAT st p0 >= 6 & len f = n & len g = n & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((insert-sort (n,(p0 + 1))),P,s),p0 holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) let f, g be FinSequence of INT ; ::_thesis: for p0, n being Element of NAT st p0 >= 6 & len f = n & len g = n & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((insert-sort (n,(p0 + 1))),P,s),p0 holds ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) let p0, n be Element of NAT ; ::_thesis: ( p0 >= 6 & len f = n & len g = n & f is_FinSequence_on s,p0 & g is_FinSequence_on IExec ((insert-sort (n,(p0 + 1))),P,s),p0 implies ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) ) assume that A1: p0 >= 6 and A2: ( len f = n & len g = n ) and A3: f is_FinSequence_on s,p0 and A4: g is_FinSequence_on IExec ((insert-sort (n,(p0 + 1))),P,s),p0 ; ::_thesis: ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) A5: p0 + 1 >= 6 + 1 by A1, XREAL_1:6; set i1 = GBP := 0; set i2 = (GBP,1) := 0; set i3 = (GBP,2) := (n - 1); set i4 = (GBP,3) := (p0 + 1); set I4 = (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1)); set t1 = IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s); set t2 = IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s); set t3 = IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s); set t4 = Exec ((GBP := 0),s); A6: (Exec ((GBP := 0),s)) . GBP = 0 by SCMPDS_2:45; then A7: DataLoc (((Exec ((GBP := 0),s)) . GBP),1) = intpos (0 + 1) by SCMP_GCD:1; A8: (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . GBP = (Exec (((GBP,1) := 0),(Exec ((GBP := 0),s)))) . GBP by SCMPDS_5:42 .= 0 by A6, A7, AMI_3:10, SCMPDS_2:46 ; then A9: DataLoc (((IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . GBP),2) = intpos (0 + 2) by SCMP_GCD:1; A10: (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . GBP = (Exec (((GBP,2) := (n - 1)),(IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A8, A9, AMI_3:10, SCMPDS_2:46 ; then A11: DataLoc (((IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; A12: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>_3_holds_ (IExec_(((GBP_:=_0)_';'_((GBP,1)_:=_0)),P,s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i > 3 implies (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . (intpos i) = s . (intpos i) ) assume A13: i > 3 ; ::_thesis: (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . (intpos i) = s . (intpos i) then A14: i > 1 by XXREAL_0:2; thus (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . (intpos i) = (Exec (((GBP,1) := 0),(Exec ((GBP := 0),s)))) . (intpos i) by SCMPDS_5:42 .= (Exec ((GBP := 0),s)) . (intpos i) by A7, A14, AMI_3:10, SCMPDS_2:46 .= s . (intpos i) by A13, AMI_3:10, SCMPDS_2:45 ; ::_thesis: verum end; A15: now__::_thesis:_for_i_being_Element_of_NAT_st_i_>_3_holds_ (IExec_((((GBP_:=_0)_';'_((GBP,1)_:=_0))_';'_((GBP,2)_:=_(n_-_1))),P,s))_._(intpos_i)_=_s_._(intpos_i) let i be Element of NAT ; ::_thesis: ( i > 3 implies (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . (intpos i) = s . (intpos i) ) assume A16: i > 3 ; ::_thesis: (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . (intpos i) = s . (intpos i) then A17: i > 2 by XXREAL_0:2; thus (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . (intpos i) = (Exec (((GBP,2) := (n - 1)),(IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)))) . (intpos i) by SCMPDS_5:41 .= (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . (intpos i) by A9, A17, AMI_3:10, SCMPDS_2:46 .= s . (intpos i) by A12, A16 ; ::_thesis: verum end; now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_f_holds_ (Initialize_(IExec_(((((GBP_:=_0)_';'_((GBP,1)_:=_0))_';'_((GBP,2)_:=_(n_-_1)))_';'_((GBP,3)_:=_(p0_+_1))),P,s)))_._(intpos_(p0_+_i))_=_f_._i let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len f implies (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . (intpos (p0 + i)) = f . i ) assume that A18: 1 <= i and A19: i <= len f ; ::_thesis: (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . (intpos (p0 + i)) = f . i set pi = p0 + i; p0 + i >= p0 + 1 by A18, XREAL_1:6; then p0 + i >= 7 by A5, XXREAL_0:2; then A20: p0 + i > 3 by XXREAL_0:2; thus (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . (intpos (p0 + i)) = (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos (p0 + i)) by SCMPDS_5:15 .= (Exec (((GBP,3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)))) . (intpos (p0 + i)) by SCMPDS_5:41 .= (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . (intpos (p0 + i)) by A11, A20, AMI_3:10, SCMPDS_2:46 .= s . (intpos (p0 + i)) by A15, A20 .= f . i by A3, A18, A19, Def1 ; ::_thesis: verum end; then A21: f is_FinSequence_on Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)),p0 by Def1; A22: (IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)) . (intpos 1) = (Exec (((GBP,1) := 0),(Exec ((GBP := 0),s)))) . (intpos 1) by SCMPDS_5:42 .= 0 by A7, SCMPDS_2:46 ; A23: (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . (intpos 1) = (Exec (((GBP,2) := (n - 1)),(IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)))) . (intpos 1) by SCMPDS_5:41 .= 0 by A22, A9, AMI_3:10, SCMPDS_2:46 ; A24: ( (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1)) is_closed_on s,P & (((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1)) is_halting_on s,P ) by SCMPDS_6:20, SCMPDS_6:21; A25: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . GBP = (Exec (((GBP,3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A10, A11, AMI_3:10, SCMPDS_2:46 ; A26: (IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)) . (intpos 2) = (Exec (((GBP,2) := (n - 1)),(IExec (((GBP := 0) ';' ((GBP,1) := 0)),P,s)))) . (intpos 2) by SCMPDS_5:41 .= n - 1 by A9, SCMPDS_2:46 ; A27: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 3) = (Exec (((GBP,3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= p0 + 1 by A11, SCMPDS_2:46 ; A28: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 3) = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . (intpos 3) by SCMPDS_5:15; A29: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 1) = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . (intpos 1) by SCMPDS_5:15; A30: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . GBP = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . GBP by SCMPDS_5:15; A31: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 2) = (Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))) . (intpos 2) by SCMPDS_5:15; A32: (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 1) = (Exec (((GBP,3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= 0 by A23, A11, AMI_3:10, SCMPDS_2:46 ; then (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 3) >= ((IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 1)) + 7 by A27, A5; then ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)),P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)),P ) by A25, Lm12, A28, A29, A30; then A33: ( for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_closed_on IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s),P & for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))) is_halting_on IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s),P ) by SCMPDS_6:125, SCMPDS_6:126; now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_g_holds_ g_._i_=_(IExec_((for-down_(GBP,2,1,(((((AddTo_(GBP,3,1))_';'_((GBP,4)_:=_(GBP,3)))_';'_(AddTo_(GBP,1,1)))_';'_((GBP,6)_:=_(GBP,1)))_';'_(while>0_(GBP,6,((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(SubFrom_(GBP,5,(intpos_4),0)))_';'_(if>0_(GBP,5,((((((GBP,5)_:=_((intpos_4),(-_1)))_';'_(((intpos_4),(-_1))_:=_((intpos_4),0)))_';'_(((intpos_4),0)_:=_(GBP,5)))_';'_(AddTo_(GBP,4,(-_1))))_';'_(AddTo_(GBP,6,(-_1)))),(Load_((GBP,6)_:=_0)))))))))),P,(Initialize_(IExec_(((((GBP_:=_0)_';'_((GBP,1)_:=_0))_';'_((GBP,2)_:=_(n_-_1)))_';'_((GBP,3)_:=_(p0_+_1))),P,s)))))_._(intpos_(p0_+_i)) let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len g implies g . i = (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))))) . (intpos (p0 + i)) ) assume ( 1 <= i & i <= len g ) ; ::_thesis: g . i = (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))))) . (intpos (p0 + i)) hence g . i = (IExec ((((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))) ';' (for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0))))))))))),P,s)) . (intpos (p0 + i)) by A4, Def1 .= (IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s))))) . (intpos (p0 + i)) by A24, A33, SCMPDS_7:30 ; ::_thesis: verum end; then A34: g is_FinSequence_on IExec ((for-down (GBP,2,1,(((((AddTo (GBP,3,1)) ';' ((GBP,4) := (GBP,3))) ';' (AddTo (GBP,1,1))) ';' ((GBP,6) := (GBP,1))) ';' (while>0 (GBP,6,((((GBP,5) := ((intpos 4),(- 1))) ';' (SubFrom (GBP,5,(intpos 4),0))) ';' (if>0 (GBP,5,((((((GBP,5) := ((intpos 4),(- 1))) ';' (((intpos 4),(- 1)) := ((intpos 4),0))) ';' (((intpos 4),0) := (GBP,5))) ';' (AddTo (GBP,4,(- 1)))) ';' (AddTo (GBP,6,(- 1)))),(Load ((GBP,6) := 0)))))))))),P,(Initialize (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)))),p0 by Def1; (IExec (((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))) ';' ((GBP,3) := (p0 + 1))),P,s)) . (intpos 2) = (Exec (((GBP,3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((GBP,1) := 0)) ';' ((GBP,2) := (n - 1))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= n - 1 by A26, A11, AMI_3:10, SCMPDS_2:46 ; hence ( f,g are_fiberwise_equipotent & g is_non_decreasing_on 1,n ) by A1, A2, A25, A32, A27, A21, A34, Lm15, A28, A29, A30, A31; ::_thesis: verum end;