:: SCPINVAR semantic presentation begin theorem Th5: :: SCPINVAR:1 for i, j being Instruction of SCMPDS for I being Program of SCMPDS holds ( ((i ';' j) ';' I) . 0 = i & ((i ';' j) ';' I) . 1 = j ) proof let i, j be Instruction of SCMPDS; ::_thesis: for I being Program of SCMPDS holds ( ((i ';' j) ';' I) . 0 = i & ((i ';' j) ';' I) . 1 = j ) let I be Program of SCMPDS; ::_thesis: ( ((i ';' j) ';' I) . 0 = i & ((i ';' j) ';' I) . 1 = j ) set jI = j ';' I; A1: (i ';' j) ';' I = i ';' (j ';' I) by SCMPDS_4:16 .= (Load i) ';' (j ';' I) by SCMPDS_4:def_2 ; 0 in dom (Load i) by COMPOS_1:50; hence ((i ';' j) ';' I) . 0 = (Load i) . 0 by A1, AFINSQ_1:def_3 .= i by COMPOS_1:52 ; ::_thesis: ((i ';' j) ';' I) . 1 = j A2: 0 in dom (Load j) by COMPOS_1:50; card (j ';' I) = (card I) + 1 by SCMPDS_6:6; then 0 < card (j ';' I) by NAT_1:5; then A3: ( card (Load i) = 1 & 0 in dom (j ';' I) ) by AFINSQ_1:66, COMPOS_1:54; thus ((i ';' j) ';' I) . 1 = ((Load i) ';' (j ';' I)) . (0 + 1) by A1 .= (j ';' I) . 0 by A3, AFINSQ_1:def_3 .= ((Load j) ';' I) . 0 by SCMPDS_4:def_2 .= (Load j) . 0 by A2, AFINSQ_1:def_3 .= j by COMPOS_1:52 ; ::_thesis: verum end; theorem Th6: :: SCPINVAR:2 for a, b being Int_position ex f being Function of (product (the_Values_of SCMPDS)),NAT st for s being State of SCMPDS holds ( ( s . a = s . b implies f . s = 0 ) & ( s . a <> s . b implies f . s = max ((abs (s . a)),(abs (s . b))) ) ) proof let a, b be Int_position; ::_thesis: ex f being Function of (product (the_Values_of SCMPDS)),NAT st for s being State of SCMPDS holds ( ( s . a = s . b implies f . s = 0 ) & ( s . a <> s . b implies f . s = max ((abs (s . a)),(abs (s . b))) ) ) defpred S1[ set , set ] means ex t being State of SCMPDS st ( t = $1 & ( t . a = t . b implies $2 = 0 ) & ( t . a <> t . b implies $2 = max ((abs (t . a)),(abs (t . b))) ) ); A1: now__::_thesis:_for_s_being_Element_of_product_(the_Values_of_SCMPDS)_ex_y_being_Element_of_NAT_st_S1[s,y] let s be Element of product (the_Values_of SCMPDS); ::_thesis: ex y being Element of NAT st S1[y,b2] percases ( s . a = s . b or s . a <> s . b ) ; supposeA2: s . a = s . b ; ::_thesis: ex y being Element of NAT st S1[y,b2] reconsider y = 0 as Element of NAT ; take y = y; ::_thesis: S1[s,y] thus S1[s,y] by A2; ::_thesis: verum end; supposeA3: s . a <> s . b ; ::_thesis: ex y being Element of NAT st S1[y,b2] set mm = max ((abs (s . a)),(abs (s . b))); reconsider y = max ((abs (s . a)),(abs (s . b))) as Element of NAT by XXREAL_0:16; take y = y; ::_thesis: S1[s,y] thus S1[s,y] by A3; ::_thesis: verum end; end; end; consider f being Function of (product (the_Values_of SCMPDS)),NAT such that A4: for s being Element of product (the_Values_of SCMPDS) holds S1[s,f . s] from FUNCT_2:sch_3(A1); A5: for s being State of SCMPDS holds S1[s,f . s] proof let s be State of SCMPDS; ::_thesis: S1[s,f . s] reconsider s = s as Element of product (the_Values_of SCMPDS) by CARD_3:107; S1[s,f . s] by A4; hence S1[s,f . s] ; ::_thesis: verum end; take f ; ::_thesis: for s being State of SCMPDS holds ( ( s . a = s . b implies f . s = 0 ) & ( s . a <> s . b implies f . s = max ((abs (s . a)),(abs (s . b))) ) ) hereby ::_thesis: verum let s be State of SCMPDS; ::_thesis: ( ( s . a = s . b implies f . s = 0 ) & ( s . a <> s . b implies f . s = max ((abs (s . a)),(abs (s . b))) ) ) S1[s,f . s] by A5; hence ( ( s . a = s . b implies f . s = 0 ) & ( s . a <> s . b implies f . s = max ((abs (s . a)),(abs (s . b))) ) ) ; ::_thesis: verum end; end; theorem Th7: :: SCPINVAR:3 for a being Int_position ex f being Function of (product (the_Values_of SCMPDS)),NAT st for s being State of SCMPDS holds ( ( s . a >= 0 implies f . s = 0 ) & ( s . a < 0 implies f . s = - (s . a) ) ) proof let a be Int_position; ::_thesis: ex f being Function of (product (the_Values_of SCMPDS)),NAT st for s being State of SCMPDS holds ( ( s . a >= 0 implies f . s = 0 ) & ( s . a < 0 implies f . s = - (s . a) ) ) defpred S1[ set , set ] means ex t being State of SCMPDS st ( t = $1 & ( t . a >= 0 implies $2 = 0 ) & ( t . a < 0 implies $2 = - (t . a) ) ); A1: now__::_thesis:_for_s_being_Element_of_product_(the_Values_of_SCMPDS)_ex_y_being_Element_of_NAT_st_S1[s,y] let s be Element of product (the_Values_of SCMPDS); ::_thesis: ex y being Element of NAT st S1[y,b2] percases ( s . a >= 0 or s . a < 0 ) ; supposeA2: s . a >= 0 ; ::_thesis: ex y being Element of NAT st S1[y,b2] reconsider y = 0 as Element of NAT ; take y = y; ::_thesis: S1[s,y] thus S1[s,y] by A2; ::_thesis: verum end; supposeA3: s . a < 0 ; ::_thesis: ex y being Element of NAT st S1[y,b2] then - (s . a) > 0 by XREAL_1:58; then reconsider y = - (s . a) as Element of NAT by INT_1:3; take y = y; ::_thesis: S1[s,y] thus S1[s,y] by A3; ::_thesis: verum end; end; end; consider f being Function of (product (the_Values_of SCMPDS)),NAT such that A4: for s being Element of product (the_Values_of SCMPDS) holds S1[s,f . s] from FUNCT_2:sch_3(A1); A5: for s being State of SCMPDS holds S1[s,f . s] proof let s be State of SCMPDS; ::_thesis: S1[s,f . s] reconsider s = s as Element of product (the_Values_of SCMPDS) by CARD_3:107; S1[s,f . s] by A4; hence S1[s,f . s] ; ::_thesis: verum end; take f ; ::_thesis: for s being State of SCMPDS holds ( ( s . a >= 0 implies f . s = 0 ) & ( s . a < 0 implies f . s = - (s . a) ) ) hereby ::_thesis: verum let s be State of SCMPDS; ::_thesis: ( ( s . a >= 0 implies f . s = 0 ) & ( s . a < 0 implies f . s = - (s . a) ) ) S1[s,f . s] by A5; hence ( ( s . a >= 0 implies f . s = 0 ) & ( s . a < 0 implies f . s = - (s . a) ) ) ; ::_thesis: verum end; end; set A = NAT ; set D = SCM-Data-Loc ; begin scheme :: SCPINVAR:sch 1 WhileLEnd{ F1( State of SCMPDS) -> Element of NAT , F2() -> 0 -started State of SCMPDS, F3() -> Instruction-Sequence of SCMPDS, F4() -> halt-free shiftable Program of SCMPDS, F5() -> Int_position, F6() -> Integer, P1[ set ] } : ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),F3(),F2())))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),F3(),F2()))] ) provided A2: for t being 0 -started State of SCMPDS st P1[t] holds ( F1(t) = 0 iff t . (DataLoc ((F2() . F5()),F6())) >= 0 ) and A3: P1[F2()] and A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = F2() . F5() & t . (DataLoc ((F2() . F5()),F6())) < 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & F1((Initialize (IExec (F4(),Q,t)))) < F1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) proof set b = DataLoc ((F2() . F5()),F6()); set WHL = while<0 (F5(),F6(),F4()); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st F1(t) <= $1 & t . F5() = F2() . F5() & P1[t] holds ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t))] ); A5: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1(t) <= 0 & t . F5() = F2() . F5() & P1[t] holds ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t))] ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(t) <= 0 & t . F5() = F2() . F5() & P1[t] implies ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t))] ) ) T: Initialize t = t by MEMSTR_0:44; assume that A6: F1(t) <= 0 and A7: t . F5() = F2() . F5() and A8: P1[t] ; ::_thesis: ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t))] ) F1(t) >= 0 by NAT_1:2; then A9: F1(t) = 0 by A6, XXREAL_0:1; then t . (DataLoc ((F2() . F5()),F6())) >= 0 by A2, A8; then for b being Int_position holds (IExec ((while<0 (F5(),F6(),F4())),Q,t)) . b = t . b by A7, SCMPDS_8:12; hence ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),Q,t))] ) by A8, A9, T, SCPISORT:4; ::_thesis: verum end; A10: 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 A11: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_u_being_0_-started_State_of_SCMPDS for_U_being_Instruction-Sequence_of_SCMPDS_st_F1(u)_<=_k_+_1_&_u_._F5()_=_F2()_._F5()_&_P1[u]_holds_ (_F1((Initialize_(IExec_((while<0_(F5(),F6(),F4())),U,u))))_=_0_&_P1[_Initialize_(IExec_((while<0_(F5(),F6(),F4())),U,u))]_) let u be 0 -started State of SCMPDS; ::_thesis: for U being Instruction-Sequence of SCMPDS st F1(u) <= k + 1 & u . F5() = F2() . F5() & P1[u] holds ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),b3,b2)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),b3,b2))] ) let U be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(u) <= k + 1 & u . F5() = F2() . F5() & P1[u] implies ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1))] ) ) assume that A12: F1(u) <= k + 1 and A13: u . F5() = F2() . F5() and A14: P1[u] ; ::_thesis: ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1))] ) percases ( F1(u) = 0 or F1(u) <> 0 ) ; suppose F1(u) = 0 ; ::_thesis: ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1))] ) hence ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),U,u)))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),U,u))] ) by A5, A13, A14; ::_thesis: verum end; supposeA15: F1(u) <> 0 ; ::_thesis: ( H1( Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),b2,b1))] ) set Iu = IExec (F4(),U,u); A16: u . (DataLoc ((F2() . F5()),F6())) < 0 by A2, A14, A15; then A17: ( (IExec (F4(),U,u)) . F5() = F2() . F5() & P1[ Initialize (IExec (F4(),U,u))] ) by A4, A13, A14; deffunc H1( State of SCMPDS) -> Element of NAT = F1($1); A18: P1[u] by A14; A19: for t being 0 -started State of SCMPDS st P1[t] & H1(t) = 0 holds t . (DataLoc ((u . F5()),F6())) >= 0 by A2, A13; H1( Initialize (IExec (F4(),U,u))) < H1(u) by A4, A13, A14, A16; then H1( Initialize (IExec (F4(),U,u))) + 1 <= H1(u) by INT_1:7; then H1( Initialize (IExec (F4(),U,u))) + 1 <= k + 1 by A12, XXREAL_0:2; then A20: H1( Initialize (IExec (F4(),U,u))) <= k by XREAL_1:6; A21: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = u . F5() & t . (DataLoc ((u . F5()),F6())) < 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & H1( Initialize (IExec (F4(),Q,t))) < H1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) by A4, A13; A22: u . (DataLoc ((u . F5()),F6())) < 0 by A2, A13, A14, A15; XX: IExec ((while<0 (F5(),F6(),F4())),U,u) = IExec ((while<0 (F5(),F6(),F4())),U,(Initialize (IExec (F4(),U,u)))) from SCMPDS_8:sch_2(A22, A19, A18, A21); (Initialize (IExec (F4(),U,u))) . F5() = (IExec (F4(),U,u)) . F5() by SCMPDS_5:15; hence ( H1( Initialize (IExec ((while<0 (F5(),F6(),F4())),U,u))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),U,u))] ) by A11, A20, A17, XX; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A5, A10); then S1[F1(F2())] ; hence ( F1((Initialize (IExec ((while<0 (F5(),F6(),F4())),F3(),F2())))) = 0 & P1[ Initialize (IExec ((while<0 (F5(),F6(),F4())),F3(),F2()))] ) by A3; ::_thesis: verum end; set a1 = intpos 1; set a2 = intpos 2; set a3 = intpos 3; begin definition let n, p0 be Element of NAT ; func sum (n,p0) -> Program of SCMPDS equals :: SCPINVAR:def 1 ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))) ';' (while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))); coherence ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))) ';' (while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))) is Program of SCMPDS ; end; :: deftheorem defines sum SCPINVAR:def_1_:_ for n, p0 being Element of NAT holds sum (n,p0) = ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))) ';' (while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))); definition let ff be Function of (product (the_Values_of SCMPDS)),NAT; let s be State of SCMPDS; :: original: . redefine funcff . s -> Element of NAT ; coherence ff . s is Element of NAT by ORDINAL1:def_12; end; theorem Th8: :: SCPINVAR:4 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, b, c being Int_position for n, i, p0 being Element of NAT for f being FinSequence of INT st f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) holds ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,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, b, c being Int_position for n, i, p0 being Element of NAT for f being FinSequence of INT st f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) holds ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,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, b, c being Int_position for n, i, p0 being Element of NAT for f being FinSequence of INT st f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) holds ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a, b, c being Int_position for n, i, p0 being Element of NAT for f being FinSequence of INT st f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) holds ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let a, b, c be Int_position; ::_thesis: for n, i, p0 being Element of NAT for f being FinSequence of INT st f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) holds ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let n, i, p0 be Element of NAT ; ::_thesis: for f being FinSequence of INT st f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) holds ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let f be FinSequence of INT ; ::_thesis: ( f is_FinSequence_on s,p0 & len f = n & s . b = 0 & s . a = 0 & s . (intpos i) = - n & s . c = p0 + 1 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) implies ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) ) set Iw = IExec ((while<0 (a,i,I)),P,s); set Dw = Initialize (IExec ((while<0 (a,i,I)),P,s)); set da = DataLoc ((s . a),i); defpred S1[ State of SCMPDS] means ( ( for i being Element of NAT st i > p0 holds $1 . (intpos i) = s . (intpos i) ) & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ($1 . (intpos i)) + n & $1 . b = Sum g & $1 . (intpos i) <= 0 & $1 . c = (p0 + 1) + (len g) ) ); assume that A2: f is_FinSequence_on s,p0 and A3: len f = n and A4: s . b = 0 and A5: s . a = 0 and A6: s . (intpos i) = - n and A7: s . c = p0 + 1 ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) & not ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ) or ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) ) consider ff being Function of (product (the_Values_of SCMPDS)),NAT such that A8: for t being State of SCMPDS holds ( ( t . (DataLoc ((s . a),i)) >= 0 implies ff . t = 0 ) & ( t . (DataLoc ((s . a),i)) < 0 implies ff . t = - (t . (DataLoc ((s . a),i))) ) ) by Th7; deffunc H1( State of SCMPDS) -> Element of NAT = ff . $1; assume A9: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos i)) + n & t . b = Sum g & t . c = (p0 + 1) + (len g) ) & t . a = 0 & t . (intpos i) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos i)) + n) + 1 & (IExec (I,Q,t)) . c = (p0 + 1) + (len g) & (IExec (I,Q,t)) . b = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec (I,Q,t)) . (intpos i) = s . (intpos i) ) ) ; ::_thesis: ( (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f & while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) A10: for t being 0 -started State of SCMPDS st S1[t] holds ( ( H1(t) = 0 implies not t . (DataLoc ((s . a),i)) < 0 ) & ( t . (DataLoc ((s . a),i)) >= 0 implies H1(t) = 0 ) ) proof let t be 0 -started State of SCMPDS; ::_thesis: ( S1[t] implies ( ( H1(t) = 0 implies not t . (DataLoc ((s . a),i)) < 0 ) & ( t . (DataLoc ((s . a),i)) >= 0 implies H1(t) = 0 ) ) ) assume S1[t] ; ::_thesis: ( ( H1(t) = 0 implies not t . (DataLoc ((s . a),i)) < 0 ) & ( t . (DataLoc ((s . a),i)) >= 0 implies H1(t) = 0 ) ) hereby ::_thesis: ( t . (DataLoc ((s . a),i)) >= 0 implies H1(t) = 0 ) assume A11: H1(t) = 0 ; ::_thesis: not t . (DataLoc ((s . a),i)) < 0 assume A12: t . (DataLoc ((s . a),i)) < 0 ; ::_thesis: contradiction then t . (DataLoc ((s . a),i)) < 0 ; then H1(t) = - (t . (DataLoc ((s . a),i))) by A8; hence contradiction by A11, A12; ::_thesis: verum end; assume t . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: H1(t) = 0 then t . (DataLoc ((s . a),i)) >= 0 ; hence H1(t) = 0 by A8; ::_thesis: verum end; A13: 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,Q,t))_._a_=_t_._a_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_H1(_Initialize_(IExec_(I,Q,t)))_<_H1(t)_&_S1[_Initialize_(IExec_(I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) ) assume that A14: S1[t] and A15: t . a = s . a and A16: t . (DataLoc ((s . a),i)) < 0 ; ::_thesis: ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) consider h being FinSequence of INT such that A17: h is_FinSequence_on s,p0 and A18: ( len h = (t . (intpos i)) + n & t . b = Sum h ) and A19: t . c = (p0 + 1) + (len h) by A14; A20: t . c = (p0 + 1) + (len h) by A19; set It = IExec (I,Q,t); set Dit = Initialize (IExec (I,Q,t)); A21: for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) by A14; A23: intpos (0 + i) = DataLoc ((s . a),i) by A5, SCMP_GCD:1; A24: ( len h = (t . (intpos i)) + n & t . b = Sum h ) by A18; hence (IExec (I,Q,t)) . a = t . a by A5, A9, A15, A16, A21, A17, A20, A23; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) consider g being FinSequence of INT such that A25: g is_FinSequence_on s,p0 and A26: len g = ((t . (intpos i)) + n) + 1 and A27: (IExec (I,Q,t)) . c = (p0 + 1) + (len g) and A28: (IExec (I,Q,t)) . b = Sum g by A5, A9, A15, A16, A21, A17, A24, A20, A23; thus ( I is_closed_on t,Q & I is_halting_on t,Q ) by A5, A9, A15, A16, A21, A17, A24, A20, A23; ::_thesis: ( H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) A29: (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) + 1 by A5, A9, A15, A16, A21, A17, A24, A20, A23; hereby ::_thesis: S1[ Initialize (IExec (I,Q,t))] percases ( (IExec (I,Q,t)) . (intpos i) >= 0 or (IExec (I,Q,t)) . (intpos i) < 0 ) ; suppose (IExec (I,Q,t)) . (intpos i) >= 0 ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) then (Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i)) >= 0 by A23, SCMPDS_5:15; then A30: H1( Initialize (IExec (I,Q,t))) = 0 by A8; H1(t) <> 0 by A10, A14, A16; hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A30, NAT_1:3; ::_thesis: verum end; supposeA31: (IExec (I,Q,t)) . (intpos i) < 0 ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) t . (DataLoc ((s . a),i)) < 0 by A16; then A32: H1(t) = - (t . (DataLoc ((s . a),i))) by A8 .= - (t . (intpos i)) by A23 ; (Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i)) < 0 by A23, A31, SCMPDS_5:15; then H1( Initialize (IExec (I,Q,t))) = - ((Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i))) by A8 .= - ((t . (intpos i)) + 1) by A23, A29, SCMPDS_5:15 .= (- (t . (intpos i))) - 1 ; hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A32, XREAL_1:146; ::_thesis: verum end; end; end; thus S1[ Initialize (IExec (I,Q,t))] ::_thesis: verum proof hereby ::_thesis: ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((Initialize (IExec (I,Q,t))) . (intpos i)) + n & (Initialize (IExec (I,Q,t))) . b = Sum g & (Initialize (IExec (I,Q,t))) . (intpos i) <= 0 & (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) ) let i be Element of NAT ; ::_thesis: ( i > p0 implies (Initialize (IExec (I,Q,t))) . (intpos i) = s . (intpos i) ) assume A33: i > p0 ; ::_thesis: (Initialize (IExec (I,Q,t))) . (intpos i) = s . (intpos i) thus (Initialize (IExec (I,Q,t))) . (intpos i) = (IExec (I,Q,t)) . (intpos i) by SCMPDS_5:15 .= s . (intpos i) by A5, A9, A15, A16, A21, A17, A24, A20, A23, A33 ; ::_thesis: verum end; take g ; ::_thesis: ( g is_FinSequence_on s,p0 & len g = ((Initialize (IExec (I,Q,t))) . (intpos i)) + n & (Initialize (IExec (I,Q,t))) . b = Sum g & (Initialize (IExec (I,Q,t))) . (intpos i) <= 0 & (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) ) thus g is_FinSequence_on s,p0 by A25; ::_thesis: ( len g = ((Initialize (IExec (I,Q,t))) . (intpos i)) + n & (Initialize (IExec (I,Q,t))) . b = Sum g & (Initialize (IExec (I,Q,t))) . (intpos i) <= 0 & (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) ) thus len g = ((IExec (I,Q,t)) . (intpos i)) + n by A29, A26 .= ((Initialize (IExec (I,Q,t))) . (intpos i)) + n by SCMPDS_5:15 ; ::_thesis: ( (Initialize (IExec (I,Q,t))) . b = Sum g & (Initialize (IExec (I,Q,t))) . (intpos i) <= 0 & (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) ) thus (Initialize (IExec (I,Q,t))) . b = Sum g by A28, SCMPDS_5:15; ::_thesis: ( (Initialize (IExec (I,Q,t))) . (intpos i) <= 0 & (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) ) (Initialize (IExec (I,Q,t))) . (intpos i) = (t . (intpos i)) + 1 by A29, SCMPDS_5:15; hence (Initialize (IExec (I,Q,t))) . (intpos i) <= 0 by A16, A23, INT_1:7; ::_thesis: (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) thus (Initialize (IExec (I,Q,t))) . c = (p0 + 1) + (len g) by A27, SCMPDS_5:15; ::_thesis: verum end; end; A34: S1[s] proof thus for i being Element of NAT st i > p0 holds s . (intpos i) = s . (intpos i) ; ::_thesis: ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (s . (intpos i)) + n & s . b = Sum g & s . (intpos i) <= 0 & s . c = (p0 + 1) + (len g) ) consider h being FinSequence of INT such that A35: len h = 0 and A36: h is_FinSequence_on s,p0 by SCPISORT:2; take h ; ::_thesis: ( h is_FinSequence_on s,p0 & len h = (s . (intpos i)) + n & s . b = Sum h & s . (intpos i) <= 0 & s . c = (p0 + 1) + (len h) ) thus h is_FinSequence_on s,p0 by A36; ::_thesis: ( len h = (s . (intpos i)) + n & s . b = Sum h & s . (intpos i) <= 0 & s . c = (p0 + 1) + (len h) ) thus len h = (s . (intpos i)) + n by A6, A35 .= (s . (intpos i)) + n ; ::_thesis: ( s . b = Sum h & s . (intpos i) <= 0 & s . c = (p0 + 1) + (len h) ) h = <*> REAL by A35; hence s . b = Sum h by A4, RVSUM_1:72; ::_thesis: ( s . (intpos i) <= 0 & s . c = (p0 + 1) + (len h) ) thus s . (intpos i) <= 0 by A6; ::_thesis: s . c = (p0 + 1) + (len h) thus s . c = (p0 + 1) + (len h) by A7, A35; ::_thesis: verum end; A37: ( H1( Initialize (IExec ((while<0 (a,i,I)),P,s))) = 0 & S1[ Initialize (IExec ((while<0 (a,i,I)),P,s))] ) from SCPINVAR:sch_1(A10, A34, A13); then consider g being FinSequence of INT such that A38: g is_FinSequence_on s,p0 and A39: len g = ((Initialize (IExec ((while<0 (a,i,I)),P,s))) . (intpos i)) + n and A40: (Initialize (IExec ((while<0 (a,i,I)),P,s))) . b = Sum g and A41: (Initialize (IExec ((while<0 (a,i,I)),P,s))) . (intpos i) <= 0 ; XX: (Initialize (IExec ((while<0 (a,i,I)),P,s))) . (intpos i) = (IExec ((while<0 (a,i,I)),P,s)) . (intpos (0 + i)) by SCMPDS_5:15 .= (IExec ((while<0 (a,i,I)),P,s)) . (DataLoc ((s . a),i)) by A5, SCMP_GCD:1 ; (IExec ((while<0 (a,i,I)),P,s)) . (DataLoc ((s . a),i)) = (Initialize (IExec ((while<0 (a,i,I)),P,s))) . (DataLoc ((s . a),i)) by SCMPDS_5:15; then (Initialize (IExec ((while<0 (a,i,I)),P,s))) . (intpos i) >= 0 by A10, A37, XX; then A42: (Initialize (IExec ((while<0 (a,i,I)),P,s))) . (intpos i) = 0 by A41, XXREAL_0:1; now__::_thesis:_for_i_being_Nat_st_i_in_dom_f_holds_ f_._i_=_g_._i let i be Nat; ::_thesis: ( i in dom f implies f . i = g . i ) reconsider a = i as Element of NAT by ORDINAL1:def_12; assume i in dom f ; ::_thesis: f . i = g . i then A43: ( 1 <= i & i <= n ) by A3, FINSEQ_3:25; hence f . i = s . (intpos (p0 + a)) by A2, A3, SCPISORT:def_1 .= g . i by A38, A39, A42, A43, SCPISORT:def_1 ; ::_thesis: verum end; then f = g by A3, A39, A42, FINSEQ_2:9; hence (IExec ((while<0 (a,i,I)),P,s)) . b = Sum f by A40, SCMPDS_5:15; ::_thesis: ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) A44: for t being 0 -started State of SCMPDS st S1[t] & H1(t) = 0 holds t . (DataLoc ((s . a),i)) >= 0 by A10; ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) from SCMPDS_8:sch_1(A44, A34, A13); hence ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) ; ::_thesis: verum end; set j1 = AddTo (GBP,1,(intpos 3),0); set j2 = AddTo (GBP,2,1); set j3 = AddTo (GBP,3,1); set WB = ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)); set WH = while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))); Lm1: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for m being Element of NAT st s . GBP = 0 & s . (intpos 3) = m holds ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . GBP = 0 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (s . (intpos 1)) + (s . (intpos m)) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for m being Element of NAT st s . GBP = 0 & s . (intpos 3) = m holds ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . GBP = 0 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (s . (intpos 1)) + (s . (intpos m)) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: for m being Element of NAT st s . GBP = 0 & s . (intpos 3) = m holds ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . GBP = 0 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (s . (intpos 1)) + (s . (intpos m)) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) let m be Element of NAT ; ::_thesis: ( s . GBP = 0 & s . (intpos 3) = m implies ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . GBP = 0 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (s . (intpos 1)) + (s . (intpos m)) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) ) assume that A1: s . GBP = 0 and A2: s . (intpos 3) = m ; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . GBP = 0 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (s . (intpos 1)) + (s . (intpos m)) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) set t0 = Initialize s; set t1 = IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s)); set t2 = IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s)); set t3 = Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s)); A3: (Initialize s) . (intpos 3) = m by A2, SCMPDS_5:15; A4: (Initialize s) . GBP = 0 by A1, SCMPDS_5:15; then 0 <> abs (((Initialize s) . GBP) + 1) by ABSVALUE:def_1; then GBP <> DataLoc (((Initialize s) . GBP),1) by XTUPLE_0:1; then A5: (Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP = 0 by A4, SCMPDS_2:49; then 0 <> abs (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP) + 2) by ABSVALUE:def_1; then A6: GBP <> DataLoc (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP),2) by XTUPLE_0:1; 3 <> abs (((Initialize s) . GBP) + 1) by A4, ABSVALUE:def_1; then intpos 3 <> DataLoc (((Initialize s) . GBP),1) by XTUPLE_0:1; then A7: (Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . (intpos 3) = m by A3, SCMPDS_2:49; 3 <> abs (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP) + 2) by A5, ABSVALUE:def_1; then A8: intpos 3 <> DataLoc (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP),2) by XTUPLE_0:1; A9: (IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . (intpos 3) = (Exec ((AddTo (GBP,2,1)),(Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))))) . (intpos 3) by SCMPDS_5:42 .= m by A7, A8, SCMPDS_2:48 ; A10: (Initialize s) . (intpos 1) = s . (intpos 1) by SCMPDS_5:15; A11: DataLoc (((Initialize s) . GBP),1) = intpos (0 + 1) by A4, SCMP_GCD:1; then A12: (Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . (intpos 1) = ((Initialize s) . (intpos 1)) + ((Initialize s) . (DataLoc (((Initialize s) . (intpos 3)),0))) by SCMPDS_2:49 .= ((Initialize s) . (intpos 1)) + ((Initialize s) . (intpos (m + 0))) by A3, SCMP_GCD:1 .= (s . (intpos 1)) + (s . (intpos m)) by A10, SCMPDS_5:15 ; 1 <> abs (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP) + 2) by A5, ABSVALUE:def_1; then A13: intpos 1 <> DataLoc (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP),2) by XTUPLE_0:1; A14: DataLoc (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP),2) = intpos (0 + 2) by A5, SCMP_GCD:1; then A15: abs (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP) + 2) = 0 + 2 by XTUPLE_0:1; 2 <> abs (((Initialize s) . GBP) + 1) by A4, ABSVALUE:def_1; then intpos 2 <> DataLoc (((Initialize s) . GBP),1) by XTUPLE_0:1; then A16: (Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . (intpos 2) = (Initialize s) . (intpos 2) by SCMPDS_2:49 .= s . (intpos 2) by SCMPDS_5:15 ; A17: (IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP = (Exec ((AddTo (GBP,2,1)),(Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))))) . GBP by SCMPDS_5:42 .= 0 by A5, A6, SCMPDS_2:48 ; then A18: DataLoc (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; A19: (IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . (intpos 2) = (Exec ((AddTo (GBP,2,1)),(Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))))) . (intpos 2) by SCMPDS_5:42 .= (s . (intpos 2)) + 1 by A16, A14, SCMPDS_2:48 ; A20: (IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . (intpos 1) = (Exec ((AddTo (GBP,2,1)),(Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))))) . (intpos 1) by SCMPDS_5:42 .= (s . (intpos 1)) + (s . (intpos m)) by A12, A13, SCMPDS_2:48 ; 0 <> abs (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP) + 3) by A17, ABSVALUE:def_1; then A21: GBP <> DataLoc (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP),3) by XTUPLE_0:1; thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . GBP = (Exec ((AddTo (GBP,3,1)),(IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))))) . GBP by SCMPDS_5:41 .= 0 by A17, A21, SCMPDS_2:48 ; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (s . (intpos 1)) + (s . (intpos m)) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) 1 <> abs (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP) + 3) by A17, ABSVALUE:def_1; then A22: intpos 1 <> DataLoc (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP),3) by XTUPLE_0:1; thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 1) = (Exec ((AddTo (GBP,3,1)),(IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))))) . (intpos 1) by SCMPDS_5:41 .= (s . (intpos 1)) + (s . (intpos m)) by A20, A22, SCMPDS_2:48 ; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (s . (intpos 2)) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) 2 <> abs (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP) + 3) by A17, ABSVALUE:def_1; then A23: intpos 2 <> DataLoc (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP),3) by XTUPLE_0:1; thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 2) = (Exec ((AddTo (GBP,3,1)),(IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))))) . (intpos 2) by SCMPDS_5:41 .= (s . (intpos 2)) + 1 by A19, A23, SCMPDS_2:48 ; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = m + 1 & ( for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) ) thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos 3) = (Exec ((AddTo (GBP,3,1)),(IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))))) . (intpos 3) by SCMPDS_5:41 .= m + 1 by A9, A18, SCMPDS_2:48 ; ::_thesis: for i being Element of NAT st i > 3 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) A24: abs (((Initialize s) . GBP) + 1) = 0 + 1 by A11, XTUPLE_0:1; hereby ::_thesis: verum let i be Element of NAT ; ::_thesis: ( i > 3 implies (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) ) assume A25: i > 3 ; ::_thesis: (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = s . (intpos i) then A26: intpos i <> DataLoc (((IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . GBP),3) by A18, XTUPLE_0:1; i <> abs (((Initialize s) . GBP) + 1) by A24, A25; then A27: intpos i <> DataLoc (((Initialize s) . GBP),1) by XTUPLE_0:1; i <> abs (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP) + 2) by A15, A25; then A28: intpos i <> DataLoc (((Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . GBP),2) by XTUPLE_0:1; thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),P,(Initialize s))) . (intpos i) = (Exec ((AddTo (GBP,3,1)),(IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))))) . (intpos i) by SCMPDS_5:41 .= (IExec (((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))),P,(Initialize s))) . (intpos i) by A26, SCMPDS_2:48 .= (Exec ((AddTo (GBP,2,1)),(Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))))) . (intpos i) by SCMPDS_5:42 .= (Exec ((AddTo (GBP,1,(intpos 3),0)),(Initialize s))) . (intpos i) by A28, SCMPDS_2:48 .= (Initialize s) . (intpos i) by A27, SCMPDS_2:49 .= s . (intpos i) by SCMPDS_5:15 ; ::_thesis: verum end; end; Lm3: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n & s . (intpos 1) = 0 & s . GBP = 0 & s . (intpos 2) = - n & s . (intpos 3) = p0 + 1 holds ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n & s . (intpos 1) = 0 & s . GBP = 0 & s . (intpos 2) = - n & s . (intpos 3) = p0 + 1 holds ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n & s . (intpos 1) = 0 & s . GBP = 0 & s . (intpos 2) = - n & s . (intpos 3) = p0 + 1 holds ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) let n, p0 be Element of NAT ; ::_thesis: for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n & s . (intpos 1) = 0 & s . GBP = 0 & s . (intpos 2) = - n & s . (intpos 3) = p0 + 1 holds ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) let f be FinSequence of INT ; ::_thesis: ( p0 >= 3 & f is_FinSequence_on s,p0 & len f = n & s . (intpos 1) = 0 & s . GBP = 0 & s . (intpos 2) = - n & s . (intpos 3) = p0 + 1 implies ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) ) set a = GBP ; I: Initialize s = s by MEMSTR_0:44; assume that A1: p0 >= 3 and A2: ( f is_FinSequence_on s,p0 & len f = n & s . (intpos 1) = 0 & s . GBP = 0 & s . (intpos 2) = - n & s . (intpos 3) = p0 + 1 ) ; ::_thesis: ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_ex_g_being_FinSequence_of_INT_st_ (_g_is_FinSequence_on_s,p0_&_len_g_=_(t_._(intpos_2))_+_n_&_t_._(intpos_1)_=_Sum_g_&_t_._(intpos_3)_=_(p0_+_1)_+_(len_g)_)_&_t_._GBP_=_0_&_t_._(intpos_2)_<_0_&_(_for_i_being_Element_of_NAT_st_i_>_p0_holds_ t_._(intpos_i)_=_s_._(intpos_i)_)_holds_ (_(IExec_((((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))),Q,t))_._GBP_=_0_&_((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))_is_closed_on_t,Q_&_((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))_is_halting_on_t,Q_&_(IExec_((((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))),Q,t))_._(intpos_2)_=_(t_._(intpos_2))_+_1_&_ex_g_being_FinSequence_of_INT_st_ (_g_is_FinSequence_on_s,p0_&_len_g_=_((t_._(intpos_2))_+_n)_+_1_&_(IExec_((((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))),Q,t))_._(intpos_3)_=_(p0_+_1)_+_(len_g)_&_(IExec_((((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))),Q,t))_._(intpos_1)_=_Sum_g_)_&_(_for_i_being_Element_of_NAT_st_i_>_p0_holds_ (IExec_((((AddTo_(GBP,1,(intpos_3),0))_';'_(AddTo_(GBP,2,1)))_';'_(AddTo_(GBP,3,1))),Q,t))_._(intpos_i)_=_s_._(intpos_i)_)_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos 2)) + n & t . (intpos 1) = Sum g & t . (intpos 3) = (p0 + 1) + (len g) ) & t . GBP = 0 & t . (intpos 2) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) holds ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . GBP = 0 & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = (t . (intpos 2)) + n & t . (intpos 1) = Sum g & t . (intpos 3) = (p0 + 1) + (len g) ) & t . GBP = 0 & t . (intpos 2) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) implies ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . GBP = 0 & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) ) T: Initialize t = t by MEMSTR_0:44; given g being FinSequence of INT such that A3: g is_FinSequence_on s,p0 and A4: len g = (t . (intpos 2)) + n and A5: t . (intpos 1) = Sum g and A6: t . (intpos 3) = (p0 + 1) + (len g) ; ::_thesis: ( t . GBP = 0 & t . (intpos 2) < 0 & ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) implies ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . GBP = 0 & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) ) assume that A7: t . GBP = 0 and t . (intpos 2) < 0 ; ::_thesis: ( ( for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ) implies ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . GBP = 0 & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) ) assume A13: for i being Element of NAT st i > p0 holds t . (intpos i) = s . (intpos i) ; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . GBP = 0 & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . GBP = 0 by A6, A7, Lm1, T; ::_thesis: ( ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) thus ( ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_closed_on t,Q & ((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)) is_halting_on t,Q ) by SCMPDS_6:20, SCMPDS_6:21; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 & ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 2) = (t . (intpos 2)) + 1 by A6, A7, Lm1, T; ::_thesis: ( ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) & ( for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) ) thus ex g being FinSequence of INT st ( g is_FinSequence_on s,p0 & len g = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len g) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum g ) ::_thesis: for i being Element of NAT st i > p0 holds (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) proof consider h being FinSequence of INT such that A14: len h = (len g) + 1 and A15: h is_FinSequence_on s,p0 by SCPISORT:2; take h ; ::_thesis: ( h is_FinSequence_on s,p0 & len h = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len h) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum h ) thus h is_FinSequence_on s,p0 by A15; ::_thesis: ( len h = ((t . (intpos 2)) + n) + 1 & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len h) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum h ) thus len h = ((t . (intpos 2)) + n) + 1 by A4, A14; ::_thesis: ( (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = (p0 + 1) + (len h) & (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum h ) thus (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 3) = ((p0 + 1) + (len g)) + 1 by A6, A7, Lm1, T .= (p0 + 1) + (len h) by A14 ; ::_thesis: (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = Sum h A16: p0 + 1 > p0 by XREAL_1:29; set m = len h; A17: len h >= 1 by A14, NAT_1:11; then p0 + (len h) >= p0 + 1 by XREAL_1:6; then A18: p0 + (len h) > p0 by A16, XXREAL_0:2; reconsider q = h . (len h) as Element of INT by INT_1:def_2; A19: now__::_thesis:_for_i_being_Nat_st_1_<=_i_&_i_<=_len_h_holds_ h_._i_=_(g_^_<*q*>)_._i let i be Nat; ::_thesis: ( 1 <= i & i <= len h implies h . b1 = (g ^ <*q*>) . b1 ) A20: i in NAT by ORDINAL1:def_12; assume that A21: 1 <= i and A22: i <= len h ; ::_thesis: h . b1 = (g ^ <*q*>) . b1 percases ( i = len h or i <> len h ) ; suppose i = len h ; ::_thesis: h . b1 = (g ^ <*q*>) . b1 hence h . i = (g ^ <*q*>) . i by A14, FINSEQ_1:42; ::_thesis: verum end; suppose i <> len h ; ::_thesis: h . b1 = (g ^ <*q*>) . b1 then i < len h by A22, XXREAL_0:1; then A23: i <= len g by A14, INT_1:7; then i in Seg (len g) by A21, FINSEQ_1:1; then A24: i in dom g by FINSEQ_1:def_3; thus h . i = s . (intpos (p0 + i)) by A15, A20, A21, A22, SCPISORT:def_1 .= g . i by A3, A20, A21, A23, SCPISORT:def_1 .= (g ^ <*q*>) . i by A24, FINSEQ_1:def_7 ; ::_thesis: verum end; end; end; len (g ^ <*q*>) = len h by A14, FINSEQ_2:16; then A25: g ^ <*q*> = h by A19, FINSEQ_1:14; h . (len h) = s . (intpos (p0 + (len h))) by A15, A17, SCPISORT:def_1 .= t . (intpos ((p0 + 1) + (len g))) by A13, A14, A18 ; hence (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos 1) = (t . (intpos 1)) + (h . (len h)) by A6, A7, Lm1, T .= Sum h by A5, A25, RVSUM_1:74 ; ::_thesis: verum end; hereby ::_thesis: verum let i be Element of NAT ; ::_thesis: ( i > p0 implies (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) ) assume A28: i > p0 ; ::_thesis: (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = s . (intpos i) then i > 3 by A1, XXREAL_0:2; hence (IExec ((((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))),Q,t)) . (intpos i) = t . (intpos i) by A6, A7, Lm1, T .= s . (intpos i) by A13, A28 ; ::_thesis: verum end; end; hence ( (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize s))) . (intpos 1) = Sum f & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on s,P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on s,P ) by A2, Th8, I; ::_thesis: verum end; Lm4: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is_halting_on s,P ) let n, p0 be Element of NAT ; ::_thesis: for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is_halting_on s,P ) let f be FinSequence of INT ; ::_thesis: ( p0 >= 3 & f is_FinSequence_on s,p0 & len f = n implies ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is_halting_on s,P ) ) I: Initialize s = s by MEMSTR_0:44; assume that A1: p0 >= 3 and A2: f is_FinSequence_on s,p0 and A3: len f = n ; ::_thesis: ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is_halting_on s,P ) set a = GBP ; set i1 = GBP := 0; set i2 = (intpos 1) := 0; set i3 = (intpos 2) := (- n); set i4 = (intpos 3) := (p0 + 1); set t0 = Initialize s; set I4 = (((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1)); set t1 = IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)); set Q1 = P; set t2 = IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s)); set t3 = IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s)); set t4 = Exec ((GBP := 0),(Initialize s)); now__::_thesis:_for_i_being_Element_of_NAT_st_1_<=_i_&_i_<=_len_f_holds_ (IExec_(((((GBP_:=_0)_';'_((intpos_1)_:=_0))_';'_((intpos_2)_:=_(-_n)))_';'_((intpos_3)_:=_(p0_+_1))),P,(Initialize_s)))_._(intpos_(p0_+_i))_=_f_._i let i be Element of NAT ; ::_thesis: ( 1 <= i & i <= len f implies (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos (p0 + i)) = f . i ) assume that A4: 1 <= i and A5: i <= len f ; ::_thesis: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos (p0 + i)) = f . i A6: p0 + 1 >= 3 + 1 by A1, XREAL_1:6; A7: p0 + i >= p0 + 1 by A4, XREAL_1:6; then p0 + i <> 3 by A6, XXREAL_0:2; then A8: intpos (p0 + i) <> intpos 3 by XTUPLE_0:1; p0 + i <> 0 by A6, A7, XXREAL_0:2; then A9: intpos (p0 + i) <> GBP by XTUPLE_0:1; p0 + i <> 1 by A6, A7, XXREAL_0:2; then A10: intpos (p0 + i) <> intpos 1 by XTUPLE_0:1; p0 + i <> 2 by A6, A7, XXREAL_0:2; then A11: intpos (p0 + i) <> intpos 2 by XTUPLE_0:1; thus (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos (p0 + i)) = (Exec (((intpos 3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))))) . (intpos (p0 + i)) by SCMPDS_5:41 .= (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))) . (intpos (p0 + i)) by A8, SCMPDS_2:45 .= (Exec (((intpos 2) := (- n)),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))))) . (intpos (p0 + i)) by SCMPDS_5:41 .= (IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))) . (intpos (p0 + i)) by A11, SCMPDS_2:45 .= (Exec (((intpos 1) := 0),(Exec ((GBP := 0),(Initialize s))))) . (intpos (p0 + i)) by SCMPDS_5:42 .= (Exec ((GBP := 0),(Initialize s))) . (intpos (p0 + i)) by A10, SCMPDS_2:45 .= (Initialize s) . (intpos (p0 + i)) by A9, SCMPDS_2:45 .= s . (intpos (p0 + i)) by SCMPDS_5:15 .= f . i by A2, A4, A5, SCPISORT:def_1 ; ::_thesis: verum end; then A12: f is_FinSequence_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)),p0 by SCPISORT:def_1; A13: f is_FinSequence_on Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))),p0 proof let i be Element of NAT ; :: according to SCPISORT:def_1 ::_thesis: ( not 1 <= i or not i <= len f or f . i = (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . (intpos (p0 + i)) ) assume ( 1 <= i & i <= len f ) ; ::_thesis: f . i = (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . (intpos (p0 + i)) then f . i = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos (p0 + i)) by A12, SCPISORT:def_1; hence f . i = (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . (intpos (p0 + i)) by SCMPDS_5:15; ::_thesis: verum end; A14: (Exec ((GBP := 0),(Initialize s))) . GBP = 0 by SCMPDS_2:45; A15: (IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))) . GBP = (Exec (((intpos 1) := 0),(Exec ((GBP := 0),(Initialize s))))) . GBP by SCMPDS_5:42 .= 0 by A14, AMI_3:10, SCMPDS_2:45 ; A16: (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))) . GBP = (Exec (((intpos 2) := (- n)),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))))) . GBP by SCMPDS_5:41 .= 0 by A15, AMI_3:10, SCMPDS_2:45 ; A17: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos 3) = (Exec (((intpos 3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))))) . (intpos 3) by SCMPDS_5:41 .= p0 + 1 by SCMPDS_2:45 ; A18: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . (intpos 3) = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos 3) by SCMPDS_5:15; A19: (IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))) . (intpos 1) = (Exec (((intpos 1) := 0),(Exec ((GBP := 0),(Initialize s))))) . (intpos 1) by SCMPDS_5:42 .= 0 by SCMPDS_2:45 ; A20: (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))) . (intpos 1) = (Exec (((intpos 2) := (- n)),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))))) . (intpos 1) by SCMPDS_5:41 .= 0 by A19, AMI_3:10, SCMPDS_2:45 ; A21: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . (intpos 1) = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos 1) by SCMPDS_5:15; A22: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos 1) = (Exec (((intpos 3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))))) . (intpos 1) by SCMPDS_5:41 .= 0 by A20, AMI_3:10, SCMPDS_2:45 ; A23: (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))) . (intpos 2) = (Exec (((intpos 2) := (- n)),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,(Initialize s))))) . (intpos 2) by SCMPDS_5:41 .= - n by SCMPDS_2:45 ; A24: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . (intpos 2) = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos 2) by SCMPDS_5:15; A25: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))) . GBP = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . GBP by SCMPDS_5:15; A26: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . (intpos 2) = (Exec (((intpos 3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))))) . (intpos 2) by SCMPDS_5:41 .= - n by A23, AMI_3:10, SCMPDS_2:45 ; A27: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))) . GBP = (Exec (((intpos 3) := (p0 + 1)),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))),P,(Initialize s))))) . GBP by SCMPDS_5:41 .= 0 by A16, AMI_3:10, SCMPDS_2:45 ; then ( while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))),P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))),P ) by A1, A3, A22, A26, A17, Lm3, A13, A18, A21, A24, A25; then A28: ( while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_closed_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)),P & while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1)))) is_halting_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)),P ) by SCMPDS_6:125, SCMPDS_6:126; IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s))))) = IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))))) ; then (IExec ((while<0 (GBP,2,(((AddTo (GBP,1,(intpos 3),0)) ';' (AddTo (GBP,2,1))) ';' (AddTo (GBP,3,1))))),P,(Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := (- n))) ';' ((intpos 3) := (p0 + 1))),P,(Initialize s)))))) . (intpos 1) = Sum f by A1, A3, A27, A22, A26, A17, Lm3, A13, A18, A21, A24, A25; hence (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f by A28, I, SCPISORT:7; ::_thesis: sum (n,p0) is_halting_on s,P thus sum (n,p0) is_halting_on s,P by A28, SCPISORT:9; ::_thesis: verum end; theorem :: SCPINVAR:5 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is parahalting ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is parahalting ) let s be 0 -started State of SCMPDS; ::_thesis: for n, p0 being Element of NAT for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is parahalting ) let n, p0 be Element of NAT ; ::_thesis: for f being FinSequence of INT st p0 >= 3 & f is_FinSequence_on s,p0 & len f = n holds ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is parahalting ) let f be FinSequence of INT ; ::_thesis: ( p0 >= 3 & f is_FinSequence_on s,p0 & len f = n implies ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is parahalting ) ) assume that A1: p0 >= 3 and A2: ( f is_FinSequence_on s,p0 & len f = n ) ; ::_thesis: ( (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f & sum (n,p0) is parahalting ) thus (IExec ((sum (n,p0)),P,s)) . (intpos 1) = Sum f by A1, A2, Lm4; ::_thesis: sum (n,p0) is parahalting now__::_thesis:_for_t_being_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_holds_sum_(n,p0)_is_halting_on_t,Q let t be State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS holds sum (n,p0) is_halting_on t,Q let Q be Instruction-Sequence of SCMPDS; ::_thesis: sum (n,p0) is_halting_on t,Q consider g being FinSequence of INT such that A3: ( len g = n & g is_FinSequence_on t,p0 ) by SCPISORT:2; g is_FinSequence_on Initialize t,p0 proof let i be Element of NAT ; :: according to SCPISORT:def_1 ::_thesis: ( not 1 <= i or not i <= len g or g . i = (Initialize t) . (intpos (p0 + i)) ) assume ( 1 <= i & i <= len g ) ; ::_thesis: g . i = (Initialize t) . (intpos (p0 + i)) then g . i = t . (intpos (p0 + i)) by A3, SCPISORT:def_1; hence g . i = (Initialize t) . (intpos (p0 + i)) by SCMPDS_5:15; ::_thesis: verum end; then sum (n,p0) is_halting_on Initialize t,Q by A1, Lm4, A3; hence sum (n,p0) is_halting_on t,Q by SCMPDS_6:126; ::_thesis: verum end; hence sum (n,p0) is parahalting by SCMPDS_6:21; ::_thesis: verum end; begin scheme :: SCPINVAR:sch 2 WhileGEnd{ F1( State of SCMPDS) -> Element of NAT , F2() -> 0 -started State of SCMPDS, F3() -> Instruction-Sequence of SCMPDS, F4() -> halt-free shiftable Program of SCMPDS, F5() -> Int_position, F6() -> Integer, P1[ set ] } : ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),F3(),F2())))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),F3(),F2()))] ) provided A2: for t being 0 -started State of SCMPDS st P1[t] holds ( F1(t) = 0 iff t . (DataLoc ((F2() . F5()),F6())) <= 0 ) and A3: P1[F2()] and A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = F2() . F5() & t . (DataLoc ((F2() . F5()),F6())) > 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & F1((Initialize (IExec (F4(),Q,t)))) < F1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) proof set b = DataLoc ((F2() . F5()),F6()); set WHL = while>0 (F5(),F6(),F4()); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st F1(t) <= $1 & t . F5() = F2() . F5() & P1[t] holds ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t))] ); A5: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1(t) <= 0 & t . F5() = F2() . F5() & P1[t] holds ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t))] ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(t) <= 0 & t . F5() = F2() . F5() & P1[t] implies ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t))] ) ) T: Initialize t = t by MEMSTR_0:44; assume that A6: F1(t) <= 0 and A7: t . F5() = F2() . F5() and A8: P1[t] ; ::_thesis: ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t))] ) F1(t) >= 0 by NAT_1:2; then A9: F1(t) = 0 by A6, XXREAL_0:1; then t . (DataLoc ((F2() . F5()),F6())) <= 0 by A2, A8; then for b being Int_position holds (IExec ((while>0 (F5(),F6(),F4())),Q,t)) . b = t . b by A7, SCMPDS_8:23; hence ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),Q,t))] ) by A8, A9, T, SCPISORT:4; ::_thesis: verum end; A10: 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 A11: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_u_being_0_-started_State_of_SCMPDS for_U_being_Instruction-Sequence_of_SCMPDS_st_F1(u)_<=_k_+_1_&_u_._F5()_=_F2()_._F5()_&_P1[u]_holds_ (_F1((Initialize_(IExec_((while>0_(F5(),F6(),F4())),U,u))))_=_0_&_P1[_Initialize_(IExec_((while>0_(F5(),F6(),F4())),U,u))]_) let u be 0 -started State of SCMPDS; ::_thesis: for U being Instruction-Sequence of SCMPDS st F1(u) <= k + 1 & u . F5() = F2() . F5() & P1[u] holds ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),b3,b2)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),b3,b2))] ) let U be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(u) <= k + 1 & u . F5() = F2() . F5() & P1[u] implies ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1))] ) ) assume that A12: F1(u) <= k + 1 and A13: u . F5() = F2() . F5() and A14: P1[u] ; ::_thesis: ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1))] ) percases ( F1(u) = 0 or F1(u) <> 0 ) ; suppose F1(u) = 0 ; ::_thesis: ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1))] ) hence ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),U,u)))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),U,u))] ) by A5, A13, A14; ::_thesis: verum end; supposeA15: F1(u) <> 0 ; ::_thesis: ( H1( Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),b2,b1))] ) set Iu = IExec (F4(),U,u); A16: u . (DataLoc ((F2() . F5()),F6())) > 0 by A2, A14, A15; then A17: ( (IExec (F4(),U,u)) . F5() = F2() . F5() & P1[ Initialize (IExec (F4(),U,u))] ) by A4, A13, A14; deffunc H1( State of SCMPDS) -> Element of NAT = F1($1); A18: P1[u] by A14; A19: for t being 0 -started State of SCMPDS st P1[t] & H1(t) = 0 holds t . (DataLoc ((u . F5()),F6())) <= 0 by A2, A13; H1( Initialize (IExec (F4(),U,u))) < H1(u) by A4, A13, A14, A16; then H1( Initialize (IExec (F4(),U,u))) + 1 <= H1(u) by INT_1:7; then H1( Initialize (IExec (F4(),U,u))) + 1 <= k + 1 by A12, XXREAL_0:2; then A20: H1( Initialize (IExec (F4(),U,u))) <= k by XREAL_1:6; A21: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = u . F5() & t . (DataLoc ((u . F5()),F6())) > 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & H1( Initialize (IExec (F4(),Q,t))) < H1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) by A4, A13; A22: u . (DataLoc ((u . F5()),F6())) > 0 by A2, A13, A14, A15; XX: IExec ((while>0 (F5(),F6(),F4())),U,u) = IExec ((while>0 (F5(),F6(),F4())),U,(Initialize (IExec (F4(),U,u)))) from SCMPDS_8:sch_4(A22, A19, A18, A21); (Initialize (IExec (F4(),U,u))) . F5() = (IExec (F4(),U,u)) . F5() by SCMPDS_5:15; hence ( H1( Initialize (IExec ((while>0 (F5(),F6(),F4())),U,u))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),U,u))] ) by A11, A20, A17, XX; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A5, A10); then S1[F1(F2())] ; hence ( F1((Initialize (IExec ((while>0 (F5(),F6(),F4())),F3(),F2())))) = 0 & P1[ Initialize (IExec ((while>0 (F5(),F6(),F4())),F3(),F2()))] ) by A3; ::_thesis: verum end; begin definition let n be Element of NAT ; func Fib-macro n -> Program of SCMPDS equals :: SCPINVAR:def 2 ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)) ';' (while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))); coherence ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)) ';' (while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))) is Program of SCMPDS ; end; :: deftheorem defines Fib-macro SCPINVAR:def_2_:_ for n being Element of NAT holds Fib-macro n = ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)) ';' (while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))); theorem Th10: :: SCPINVAR:6 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, f0, f1 being Int_position for n, i being Element of NAT st s . a = 0 & s . f0 = 0 & s . f1 = 1 & s . (intpos i) = n & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) holds ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,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, f0, f1 being Int_position for n, i being Element of NAT st s . a = 0 & s . f0 = 0 & s . f1 = 1 & s . (intpos i) = n & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) holds ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,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, f0, f1 being Int_position for n, i being Element of NAT st s . a = 0 & s . f0 = 0 & s . f1 = 1 & s . (intpos i) = n & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) holds ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a, f0, f1 being Int_position for n, i being Element of NAT st s . a = 0 & s . f0 = 0 & s . f1 = 1 & s . (intpos i) = n & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) holds ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) let a, f0, f1 be Int_position; ::_thesis: for n, i being Element of NAT st s . a = 0 & s . f0 = 0 & s . f1 = 1 & s . (intpos i) = n & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) holds ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) let n, i be Element of NAT ; ::_thesis: ( s . a = 0 & s . f0 = 0 & s . f1 = 1 & s . (intpos i) = n & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) implies ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) ) set Iw = IExec ((while>0 (a,i,I)),P,s); set Dw = Initialize (IExec ((while>0 (a,i,I)),P,s)); set da = DataLoc ((s . a),i); defpred S1[ State of SCMPDS] means ( $1 . (intpos i) >= 0 & ex k being Element of NAT st ( n = ($1 . (intpos i)) + k & $1 . f0 = Fib k & $1 . f1 = Fib (k + 1) ) ); assume that A2: s . a = 0 and A3: s . f0 = 0 and A4: s . f1 = 1 and A5: s . (intpos i) = n ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS ex k being Element of NAT st ( n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 & not ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ) or ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) ) consider ff being Function of (product (the_Values_of SCMPDS)),NAT such that B6: for t being State of SCMPDS holds ( ( t . (DataLoc ((s . a),i)) <= 0 implies ff . t = 0 ) & ( t . (DataLoc ((s . a),i)) > 0 implies ff . t = t . (DataLoc ((s . a),i)) ) ) by SCMPDS_8:5; A6: for t being 0 -started State of SCMPDS holds ( ( t . (DataLoc ((s . a),i)) <= 0 implies ff . t = 0 ) & ( t . (DataLoc ((s . a),i)) > 0 implies ff . t = t . (DataLoc ((s . a),i)) ) ) by B6; deffunc H1( State of SCMPDS) -> Element of NAT = ff . $1; A7: for t being 0 -started State of SCMPDS st S1[t] holds ( ( not H1(t) = 0 or not t . (DataLoc ((s . a),i)) > 0 ) & ( t . (DataLoc ((s . a),i)) <= 0 implies H1(t) = 0 ) ) by A6; assume A9: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos i)) + k & t . f0 = Fib k & t . f1 = Fib (k + 1) & t . a = 0 & t . (intpos i) > 0 holds ( (IExec (I,Q,t)) . a = 0 & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 & (IExec (I,Q,t)) . f0 = Fib (k + 1) & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) ; ::_thesis: ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) & while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) A10: 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,Q,t))_._a_=_t_._a_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_H1(_Initialize_(IExec_(I,Q,t)))_<_H1(t)_&_S1[_Initialize_(IExec_(I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) ) assume that A11: S1[t] and A12: t . a = s . a and A13: t . (DataLoc ((s . a),i)) > 0 ; ::_thesis: ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) set It = IExec (I,Q,t); set Dit = Initialize (IExec (I,Q,t)); consider k being Element of NAT such that A14: n = (t . (intpos i)) + k and A15: t . f0 = Fib k and A16: t . f1 = Fib (k + 1) by A11; A17: t . f1 = Fib (k + 1) by A16; A18: intpos (0 + i) = DataLoc ((s . a),i) by A2, SCMP_GCD:1; A19: ( n = (t . (intpos i)) + k & t . f0 = Fib k ) by A14, A15; hence (IExec (I,Q,t)) . a = t . a by A2, A9, A12, A13, A17, A18; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) thus ( I is_closed_on t,Q & I is_halting_on t,Q ) by A2, A9, A12, A13, A19, A17, A18; ::_thesis: ( H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) A20: (IExec (I,Q,t)) . (intpos i) = (t . (intpos i)) - 1 by A2, A9, A12, A13, A19, A17, A18; hereby ::_thesis: S1[ Initialize (IExec (I,Q,t))] percases ( (IExec (I,Q,t)) . (intpos i) <= 0 or (IExec (I,Q,t)) . (intpos i) > 0 ) ; suppose (IExec (I,Q,t)) . (intpos i) <= 0 ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) then (Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i)) <= 0 by A18, SCMPDS_5:15; then A21: H1( Initialize (IExec (I,Q,t))) = 0 by A6; H1(t) <> 0 by A7, A11, A13; hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A21, NAT_1:3; ::_thesis: verum end; supposeA22: (IExec (I,Q,t)) . (intpos i) > 0 ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) t . (DataLoc ((s . a),i)) > 0 by A13; then A23: H1(t) = t . (DataLoc ((s . a),i)) by A6 .= t . (intpos i) by A18 ; (Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i)) > 0 by A18, A22, SCMPDS_5:15; then H1( Initialize (IExec (I,Q,t))) = (Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i)) by A6 .= (t . (intpos i)) - 1 by A18, A20, SCMPDS_5:15 ; hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A23, XREAL_1:146; ::_thesis: verum end; end; end; thus S1[ Initialize (IExec (I,Q,t))] ::_thesis: verum proof t . (intpos i) >= 1 + 0 by A13, A18, INT_1:7; then (t . (intpos i)) - 1 >= 0 by XREAL_1:48; hence (Initialize (IExec (I,Q,t))) . (intpos i) >= 0 by A20, SCMPDS_5:15; ::_thesis: ex k being Element of NAT st ( n = ((Initialize (IExec (I,Q,t))) . (intpos i)) + k & (Initialize (IExec (I,Q,t))) . f0 = Fib k & (Initialize (IExec (I,Q,t))) . f1 = Fib (k + 1) ) take m = k + 1; ::_thesis: ( n = ((Initialize (IExec (I,Q,t))) . (intpos i)) + m & (Initialize (IExec (I,Q,t))) . f0 = Fib m & (Initialize (IExec (I,Q,t))) . f1 = Fib (m + 1) ) thus n = (((t . (intpos i)) - 1) + 1) + k by A14 .= (((Initialize (IExec (I,Q,t))) . (intpos i)) + 1) + k by A20, SCMPDS_5:15 .= ((Initialize (IExec (I,Q,t))) . (intpos i)) + m ; ::_thesis: ( (Initialize (IExec (I,Q,t))) . f0 = Fib m & (Initialize (IExec (I,Q,t))) . f1 = Fib (m + 1) ) ( (IExec (I,Q,t)) . f0 = Fib m & (IExec (I,Q,t)) . f1 = Fib ((k + 1) + 1) ) by A2, A9, A12, A13, A19, A17, A18; hence ( (Initialize (IExec (I,Q,t))) . f0 = Fib m & (Initialize (IExec (I,Q,t))) . f1 = Fib (m + 1) ) by SCMPDS_5:15; ::_thesis: verum end; end; A24: S1[s] proof s . (intpos i) = n by A5; hence s . (intpos i) >= 0 by NAT_1:2; ::_thesis: ex k being Element of NAT st ( n = (s . (intpos i)) + k & s . f0 = Fib k & s . f1 = Fib (k + 1) ) take k = 0 ; ::_thesis: ( n = (s . (intpos i)) + k & s . f0 = Fib k & s . f1 = Fib (k + 1) ) thus n = (s . (intpos i)) + k by A5; ::_thesis: ( s . f0 = Fib k & s . f1 = Fib (k + 1) ) thus s . f0 = Fib k by A3, PRE_FF:1; ::_thesis: s . f1 = Fib (k + 1) thus s . f1 = Fib (k + 1) by A4, PRE_FF:1; ::_thesis: verum end; A25: ( H1( Initialize (IExec ((while>0 (a,i,I)),P,s))) = 0 & S1[ Initialize (IExec ((while>0 (a,i,I)),P,s))] ) from SCPINVAR:sch_2(A7, A24, A10); X1: (Initialize (IExec ((while>0 (a,i,I)),P,s))) . (DataLoc ((s . a),i)) = (IExec ((while>0 (a,i,I)),P,s)) . (DataLoc ((s . a),i)) by SCMPDS_5:15; (Initialize (IExec ((while>0 (a,i,I)),P,s))) . (intpos i) = (IExec ((while>0 (a,i,I)),P,s)) . (intpos (0 + i)) by SCMPDS_5:15 .= (IExec ((while>0 (a,i,I)),P,s)) . (DataLoc ((s . a),i)) by A2, SCMP_GCD:1 ; then (Initialize (IExec ((while>0 (a,i,I)),P,s))) . (intpos i) <= 0 by A7, A25, X1; then (Initialize (IExec ((while>0 (a,i,I)),P,s))) . (intpos i) = 0 by A25, XXREAL_0:1; hence ( (IExec ((while>0 (a,i,I)),P,s)) . f0 = Fib n & (IExec ((while>0 (a,i,I)),P,s)) . f1 = Fib (n + 1) ) by A25, SCMPDS_5:15; ::_thesis: ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) A26: for t being 0 -started State of SCMPDS st S1[t] & H1(t) = 0 holds t . (DataLoc ((s . a),i)) <= 0 by A7; ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) from SCMPDS_8:sch_3(A26, A24, A10); hence ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) ; ::_thesis: verum end; set j1 = (GBP,4) := (GBP,2); set j2 = AddTo (GBP,2,GBP,1); set j3 = (GBP,1) := (GBP,4); set j4 = AddTo (GBP,3,(- 1)); set WB = ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))); set WH = while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))); Lm5: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . GBP = 0 holds ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . GBP = 0 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 1) = s . (intpos 2) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 1)) + (s . (intpos 2)) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . GBP = 0 holds ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . GBP = 0 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 1) = s . (intpos 2) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 1)) + (s . (intpos 2)) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: ( s . GBP = 0 implies ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . GBP = 0 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 1) = s . (intpos 2) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 1)) + (s . (intpos 2)) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 ) ) set t0 = s; set t1 = IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s); set t2 = IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s); set Q0 = P; set t3 = IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s); set t4 = Exec (((GBP,4) := (GBP,2)),s); set a4 = intpos 4; assume s . GBP = 0 ; ::_thesis: ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . GBP = 0 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 1) = s . (intpos 2) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 1)) + (s . (intpos 2)) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 ) then A5: s . GBP = 0 ; then DataLoc ((s . GBP),4) = intpos (0 + 4) by SCMP_GCD:1; then A6: (Exec (((GBP,4) := (GBP,2)),s)) . (intpos 4) = s . (DataLoc ((s . GBP),2)) by SCMPDS_2:47 .= s . (intpos (0 + 2)) by A5, SCMP_GCD:1 .= s . (intpos 2) ; 0 <> abs ((s . GBP) + 4) by A5, ABSVALUE:def_1; then GBP <> DataLoc ((s . GBP),4) by XTUPLE_0:1; then A7: (Exec (((GBP,4) := (GBP,2)),s)) . GBP = 0 by A5, SCMPDS_2:47; then A8: DataLoc (((Exec (((GBP,4) := (GBP,2)),s)) . GBP),2) = intpos (0 + 2) by SCMP_GCD:1; 0 <> abs (((Exec (((GBP,4) := (GBP,2)),s)) . GBP) + 2) by A7, ABSVALUE:def_1; then A9: GBP <> DataLoc (((Exec (((GBP,4) := (GBP,2)),s)) . GBP),2) by XTUPLE_0:1; A10: (IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP = (Exec ((AddTo (GBP,2,GBP,1)),(Exec (((GBP,4) := (GBP,2)),s)))) . GBP by SCMPDS_5:42 .= 0 by A7, A9, SCMPDS_2:49 ; then A11: DataLoc (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP),1) = intpos (0 + 1) by SCMP_GCD:1; 4 <> abs (((Exec (((GBP,4) := (GBP,2)),s)) . GBP) + 2) by A7, ABSVALUE:def_1; then A12: intpos 4 <> DataLoc (((Exec (((GBP,4) := (GBP,2)),s)) . GBP),2) by XTUPLE_0:1; A13: (IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . (intpos 4) = (Exec ((AddTo (GBP,2,GBP,1)),(Exec (((GBP,4) := (GBP,2)),s)))) . (intpos 4) by SCMPDS_5:42 .= s . (intpos 2) by A6, A12, SCMPDS_2:49 ; A14: (IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . (intpos 1) = (Exec (((GBP,1) := (GBP,4)),(IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= (IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . (DataLoc (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP),4)) by A11, SCMPDS_2:47 .= s . (intpos 2) by A10, A13, SCMP_GCD:1 ; 3 <> abs (((Exec (((GBP,4) := (GBP,2)),s)) . GBP) + 2) by A7, ABSVALUE:def_1; then A15: intpos 3 <> DataLoc (((Exec (((GBP,4) := (GBP,2)),s)) . GBP),2) by XTUPLE_0:1; 2 <> abs ((s . GBP) + 4) by A5, ABSVALUE:def_1; then intpos 2 <> DataLoc ((s . GBP),4) by XTUPLE_0:1; then A16: (Exec (((GBP,4) := (GBP,2)),s)) . (intpos 2) = s . (intpos 2) by SCMPDS_2:47 .= s . (intpos 2) ; 1 <> abs ((s . GBP) + 4) by A5, ABSVALUE:def_1; then intpos 1 <> DataLoc ((s . GBP),4) by XTUPLE_0:1; then A17: (Exec (((GBP,4) := (GBP,2)),s)) . (intpos 1) = s . (intpos 1) by SCMPDS_2:47 .= s . (intpos 1) ; 3 <> abs ((s . GBP) + 4) by A5, ABSVALUE:def_1; then intpos 3 <> DataLoc ((s . GBP),4) by XTUPLE_0:1; then A18: (Exec (((GBP,4) := (GBP,2)),s)) . (intpos 3) = s . (intpos 3) by SCMPDS_2:47 .= s . (intpos 3) ; 0 <> abs (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP) + 1) by A10, ABSVALUE:def_1; then A19: GBP <> DataLoc (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP),1) by XTUPLE_0:1; A20: (IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP = (Exec (((GBP,1) := (GBP,4)),(IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A10, A19, SCMPDS_2:47 ; then A21: DataLoc (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; 2 <> abs (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP) + 1) by A10, ABSVALUE:def_1; then A22: intpos 2 <> DataLoc (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP),1) by XTUPLE_0:1; A23: (IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . (intpos 2) = (Exec ((AddTo (GBP,2,GBP,1)),(Exec (((GBP,4) := (GBP,2)),s)))) . (intpos 2) by SCMPDS_5:42 .= ((Exec (((GBP,4) := (GBP,2)),s)) . (intpos 2)) + ((Exec (((GBP,4) := (GBP,2)),s)) . (DataLoc (((Exec (((GBP,4) := (GBP,2)),s)) . GBP),1))) by A8, SCMPDS_2:49 .= (s . (intpos 2)) + (s . (intpos 1)) by A7, A17, A16, SCMP_GCD:1 ; A24: (IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . (intpos 2) = (Exec (((GBP,1) := (GBP,4)),(IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= (s . (intpos 2)) + (s . (intpos 1)) by A23, A22, SCMPDS_2:47 ; 3 <> abs (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP) + 1) by A10, ABSVALUE:def_1; then A25: intpos 3 <> DataLoc (((IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . GBP),1) by XTUPLE_0:1; 0 <> abs (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP) + 3) by A20, ABSVALUE:def_1; then A26: GBP <> DataLoc (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP),3) by XTUPLE_0:1; thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . GBP = (Exec ((AddTo (GBP,3,(- 1))),(IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A20, A26, SCMPDS_2:48 ; ::_thesis: ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 1) = s . (intpos 2) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 1)) + (s . (intpos 2)) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 ) 1 <> abs (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP) + 3) by A20, ABSVALUE:def_1; then A27: intpos 1 <> DataLoc (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP),3) by XTUPLE_0:1; thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 1) = (Exec ((AddTo (GBP,3,(- 1))),(IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= s . (intpos 2) by A14, A27, SCMPDS_2:48 ; ::_thesis: ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (s . (intpos 1)) + (s . (intpos 2)) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 ) 2 <> abs (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP) + 3) by A20, ABSVALUE:def_1; then A28: intpos 2 <> DataLoc (((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . GBP),3) by XTUPLE_0:1; A29: (IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)) . (intpos 3) = (Exec ((AddTo (GBP,2,GBP,1)),(Exec (((GBP,4) := (GBP,2)),s)))) . (intpos 3) by SCMPDS_5:42 .= s . (intpos 3) by A18, A15, SCMPDS_2:49 ; A30: (IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . (intpos 3) = (Exec (((GBP,1) := (GBP,4)),(IExec ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= s . (intpos 3) by A29, A25, SCMPDS_2:47 ; thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 2) = (Exec ((AddTo (GBP,3,(- 1))),(IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= (s . (intpos 1)) + (s . (intpos 2)) by A24, A28, SCMPDS_2:48 ; ::_thesis: (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (s . (intpos 3)) - 1 thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),P,s)) . (intpos 3) = (Exec ((AddTo (GBP,3,(- 1))),(IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= ((IExec (((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))),P,s)) . (intpos 3)) + (- 1) by A21, SCMPDS_2:48 .= (s . (intpos 3)) - 1 by A30 ; ::_thesis: verum end; Lm7: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for n being Element of NAT st s . GBP = 0 & s . (intpos 1) = 0 & s . (intpos 2) = 1 & s . (intpos 3) = n holds ( (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 1) = Fib n & (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 2) = Fib (n + 1) & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on s,P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for n being Element of NAT st s . GBP = 0 & s . (intpos 1) = 0 & s . (intpos 2) = 1 & s . (intpos 3) = n holds ( (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 1) = Fib n & (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 2) = Fib (n + 1) & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on s,P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: for n being Element of NAT st s . GBP = 0 & s . (intpos 1) = 0 & s . (intpos 2) = 1 & s . (intpos 3) = n holds ( (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 1) = Fib n & (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 2) = Fib (n + 1) & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on s,P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on s,P ) let n be Element of NAT ; ::_thesis: ( s . GBP = 0 & s . (intpos 1) = 0 & s . (intpos 2) = 1 & s . (intpos 3) = n implies ( (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 1) = Fib n & (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 2) = Fib (n + 1) & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on s,P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on s,P ) ) set a = GBP ; A1: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS for_k_being_Element_of_NAT_st_n_=_(t_._(intpos_3))_+_k_&_t_._(intpos_1)_=_Fib_k_&_t_._(intpos_2)_=_Fib_(k_+_1)_&_t_._GBP_=_0_&_t_._(intpos_3)_>_0_holds_ (_(IExec_((((((GBP,4)_:=_(GBP,2))_';'_(AddTo_(GBP,2,GBP,1)))_';'_((GBP,1)_:=_(GBP,4)))_';'_(AddTo_(GBP,3,(-_1)))),Q,t))_._GBP_=_0_&_((((GBP,4)_:=_(GBP,2))_';'_(AddTo_(GBP,2,GBP,1)))_';'_((GBP,1)_:=_(GBP,4)))_';'_(AddTo_(GBP,3,(-_1)))_is_closed_on_t,Q_&_((((GBP,4)_:=_(GBP,2))_';'_(AddTo_(GBP,2,GBP,1)))_';'_((GBP,1)_:=_(GBP,4)))_';'_(AddTo_(GBP,3,(-_1)))_is_halting_on_t,Q_&_(IExec_((((((GBP,4)_:=_(GBP,2))_';'_(AddTo_(GBP,2,GBP,1)))_';'_((GBP,1)_:=_(GBP,4)))_';'_(AddTo_(GBP,3,(-_1)))),Q,t))_._(intpos_3)_=_(t_._(intpos_3))_-_1_&_(IExec_((((((GBP,4)_:=_(GBP,2))_';'_(AddTo_(GBP,2,GBP,1)))_';'_((GBP,1)_:=_(GBP,4)))_';'_(AddTo_(GBP,3,(-_1)))),Q,t))_._(intpos_1)_=_Fib_(k_+_1)_&_(IExec_((((((GBP,4)_:=_(GBP,2))_';'_(AddTo_(GBP,2,GBP,1)))_';'_((GBP,1)_:=_(GBP,4)))_';'_(AddTo_(GBP,3,(-_1)))),Q,t))_._(intpos_2)_=_Fib_((k_+_1)_+_1)_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS for k being Element of NAT st n = (t . (intpos 3)) + k & t . (intpos 1) = Fib k & t . (intpos 2) = Fib (k + 1) & t . GBP = 0 & t . (intpos 3) > 0 holds ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . GBP = 0 & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_closed_on t,Q & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_halting_on t,Q & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: for k being Element of NAT st n = (t . (intpos 3)) + k & t . (intpos 1) = Fib k & t . (intpos 2) = Fib (k + 1) & t . GBP = 0 & t . (intpos 3) > 0 holds ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . GBP = 0 & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_closed_on t,Q & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_halting_on t,Q & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) let k be Element of NAT ; ::_thesis: ( n = (t . (intpos 3)) + k & t . (intpos 1) = Fib k & t . (intpos 2) = Fib (k + 1) & t . GBP = 0 & t . (intpos 3) > 0 implies ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . GBP = 0 & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_closed_on t,Q & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_halting_on t,Q & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) ) assume that n = (t . (intpos 3)) + k and A2: t . (intpos 1) = Fib k and A3: t . (intpos 2) = Fib (k + 1) and A4: t . GBP = 0 and t . (intpos 3) > 0 ; ::_thesis: ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . GBP = 0 & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_closed_on t,Q & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_halting_on t,Q & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . GBP = 0 by A4, Lm5; ::_thesis: ( ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_closed_on t,Q & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_halting_on t,Q & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) thus ( ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_closed_on t,Q & ((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))) is_halting_on t,Q ) by SCMPDS_6:20, SCMPDS_6:21; ::_thesis: ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 3) = (t . (intpos 3)) - 1 by A4, Lm5; ::_thesis: ( (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) & (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) ) thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 1) = Fib (k + 1) by A3, A4, Lm5; ::_thesis: (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = Fib ((k + 1) + 1) thus (IExec ((((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))),Q,t)) . (intpos 2) = (t . (intpos 1)) + (t . (intpos 2)) by A4, Lm5 .= Fib ((k + 1) + 1) by A2, A3, PRE_FF:1 ; ::_thesis: verum end; assume ( s . GBP = 0 & s . (intpos 1) = 0 & s . (intpos 2) = 1 & s . (intpos 3) = n ) ; ::_thesis: ( (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 1) = Fib n & (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 2) = Fib (n + 1) & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on s,P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on s,P ) hence ( (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 1) = Fib n & (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,s)) . (intpos 2) = Fib (n + 1) & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on s,P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on s,P ) by A1, Th10; ::_thesis: verum end; Lm8: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for n being Element of NAT holds ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for n being Element of NAT holds ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: for n being Element of NAT holds ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is_halting_on s,P ) let n be Element of NAT ; ::_thesis: ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is_halting_on s,P ) set a = GBP ; set i1 = GBP := 0; set i2 = (intpos 1) := 0; set i3 = (intpos 2) := 1; set i4 = (intpos 3) := n; set I4 = (((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n); set t1 = IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s); set t2 = IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s); set t3 = IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,s); set t4 = Exec ((GBP := 0),s); I: Initialize s = s by MEMSTR_0:44; A1: (Exec ((GBP := 0),s)) . GBP = 0 by SCMPDS_2:45; A2: (IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,s)) . GBP = (Exec (((intpos 1) := 0),(Exec ((GBP := 0),s)))) . GBP by SCMPDS_5:42 .= 0 by A1, AMI_3:10, SCMPDS_2:45 ; A3: (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)) . GBP = (Exec (((intpos 2) := 1),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A2, AMI_3:10, SCMPDS_2:45 ; A4: (IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,s)) . (intpos 1) = (Exec (((intpos 1) := 0),(Exec ((GBP := 0),s)))) . (intpos 1) by SCMPDS_5:42 .= 0 by SCMPDS_2:45 ; A5: (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)) . (intpos 1) = (Exec (((intpos 2) := 1),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,s)))) . (intpos 1) by SCMPDS_5:41 .= 0 by A4, AMI_3:10, SCMPDS_2:45 ; A6: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . (intpos 1) = (Exec (((intpos 3) := n),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)))) . (intpos 1) by SCMPDS_5:41 .= 0 by A5, AMI_3:10, SCMPDS_2:45 ; A7: (IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)) . (intpos 2) = (Exec (((intpos 2) := 1),(IExec (((GBP := 0) ';' ((intpos 1) := 0)),P,s)))) . (intpos 2) by SCMPDS_5:41 .= 1 by SCMPDS_2:45 ; A8: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . (intpos 2) = (Exec (((intpos 3) := n),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)))) . (intpos 2) by SCMPDS_5:41 .= 1 by A7, AMI_3:10, SCMPDS_2:45 ; A9: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . (intpos 3) = (Exec (((intpos 3) := n),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)))) . (intpos 3) by SCMPDS_5:41 .= n by SCMPDS_2:45 ; A10: (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . GBP = (Exec (((intpos 3) := n),(IExec ((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A3, AMI_3:10, SCMPDS_2:45 ; X1: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))) . GBP = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . GBP by SCMPDS_5:15; X2: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))) . (intpos 1) = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . (intpos 1) by SCMPDS_5:15; X3: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))) . (intpos 2) = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . (intpos 2) by SCMPDS_5:15; X4: (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))) . (intpos 3) = (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)) . (intpos 3) by SCMPDS_5:15; B11: ( while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)),P & while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)),P ) by A6, A8, A9, Lm7, A10, X1, X2, X3, X4; A11: while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s),P proof for k being Element of NAT holds IC (Comput ((P +* (stop (while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))))),(Initialize (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s)))),k)) in dom (stop (while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))))) by B11, SCMPDS_6:def_2; hence while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_closed_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s),P by SCMPDS_6:def_2; ::_thesis: verum end; A12: while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s),P proof P +* (stop (while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))))) halts_on Initialize (Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))) by B11, SCMPDS_6:def_3; hence while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1))))) is_halting_on IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s),P by SCMPDS_6:def_3; ::_thesis: verum end; (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,(Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))))) . (intpos 1) = Fib n by A10, A6, A8, A9, Lm7, X1, X2, X3, X4; hence (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n by A11, A12, SCPISORT:7; ::_thesis: ( (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is_halting_on s,P ) (IExec ((while>0 (GBP,3,(((((GBP,4) := (GBP,2)) ';' (AddTo (GBP,2,GBP,1))) ';' ((GBP,1) := (GBP,4))) ';' (AddTo (GBP,3,(- 1)))))),P,(Initialize (IExec (((((GBP := 0) ';' ((intpos 1) := 0)) ';' ((intpos 2) := 1)) ';' ((intpos 3) := n)),P,s))))) . (intpos 2) = Fib (n + 1) by A10, A6, A8, A9, Lm7, X1, X2, X3, X4; hence (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) by A11, A12, SCPISORT:7; ::_thesis: Fib-macro n is_halting_on s,P thus Fib-macro n is_halting_on s,P by A11, A12, I, SCPISORT:9; ::_thesis: verum end; theorem :: SCPINVAR:7 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for n being Element of NAT holds ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is parahalting ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS for n being Element of NAT holds ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is parahalting ) let s be 0 -started State of SCMPDS; ::_thesis: for n being Element of NAT holds ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is parahalting ) let n be Element of NAT ; ::_thesis: ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) & Fib-macro n is parahalting ) thus ( (IExec ((Fib-macro n),P,s)) . (intpos 1) = Fib n & (IExec ((Fib-macro n),P,s)) . (intpos 2) = Fib (n + 1) ) by Lm8; ::_thesis: Fib-macro n is parahalting for t being State of SCMPDS for Q being Instruction-Sequence of SCMPDS holds Fib-macro n is_halting_on t,Q proof let t be State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS holds Fib-macro n is_halting_on t,Q let Q be Instruction-Sequence of SCMPDS; ::_thesis: Fib-macro n is_halting_on t,Q Fib-macro n is_halting_on Initialize t,Q by Lm8; hence Fib-macro n is_halting_on t,Q by SCMPDS_6:126; ::_thesis: verum end; hence Fib-macro n is parahalting by SCMPDS_6:21; ::_thesis: verum end; begin definition let a be Int_position; let i be Integer; let I be Program of SCMPDS; func while<>0 (a,i,I) -> Program of SCMPDS equals :: SCPINVAR:def 3 ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I) ';' (goto (- ((card I) + 2))); coherence ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I) ';' (goto (- ((card I) + 2))) is Program of SCMPDS ; end; :: deftheorem defines while<>0 SCPINVAR:def_3_:_ for a being Int_position for i being Integer for I being Program of SCMPDS holds while<>0 (a,i,I) = ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I) ';' (goto (- ((card I) + 2))); begin theorem Th12: :: SCPINVAR:8 for a being Int_position for i being Integer for I being Program of SCMPDS holds card (while<>0 (a,i,I)) = (card I) + 3 proof let a be Int_position; ::_thesis: for i being Integer for I being Program of SCMPDS holds card (while<>0 (a,i,I)) = (card I) + 3 let i be Integer; ::_thesis: for I being Program of SCMPDS holds card (while<>0 (a,i,I)) = (card I) + 3 let I be Program of SCMPDS; ::_thesis: card (while<>0 (a,i,I)) = (card I) + 3 set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set I4 = (((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I; thus card (while<>0 (a,i,I)) = (card ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I)) + 1 by SCMP_GCD:4 .= ((card (((a,i) <>0_goto 2) ';' (goto ((card I) + 2)))) + (card I)) + 1 by AFINSQ_1:17 .= (2 + (card I)) + 1 by SCMP_GCD:5 .= (card I) + 3 ; ::_thesis: verum end; Lm9: for a being Int_position for i being Integer for I being Program of SCMPDS holds card (stop (while<>0 (a,i,I))) = (card I) + 4 proof let a be Int_position; ::_thesis: for i being Integer for I being Program of SCMPDS holds card (stop (while<>0 (a,i,I))) = (card I) + 4 let i be Integer; ::_thesis: for I being Program of SCMPDS holds card (stop (while<>0 (a,i,I))) = (card I) + 4 let I be Program of SCMPDS; ::_thesis: card (stop (while<>0 (a,i,I))) = (card I) + 4 thus card (stop (while<>0 (a,i,I))) = (card (while<>0 (a,i,I))) + 1 by COMPOS_1:55 .= ((card I) + 3) + 1 by Th12 .= (card I) + 4 ; ::_thesis: verum end; theorem Th13: :: SCPINVAR:9 for a being Int_position for i being Integer for m being Element of NAT for I being Program of SCMPDS holds ( m < (card I) + 3 iff m in dom (while<>0 (a,i,I)) ) proof let a be Int_position; ::_thesis: for i being Integer for m being Element of NAT for I being Program of SCMPDS holds ( m < (card I) + 3 iff m in dom (while<>0 (a,i,I)) ) let i be Integer; ::_thesis: for m being Element of NAT for I being Program of SCMPDS holds ( m < (card I) + 3 iff m in dom (while<>0 (a,i,I)) ) let m be Element of NAT ; ::_thesis: for I being Program of SCMPDS holds ( m < (card I) + 3 iff m in dom (while<>0 (a,i,I)) ) let I be Program of SCMPDS; ::_thesis: ( m < (card I) + 3 iff m in dom (while<>0 (a,i,I)) ) card (while<>0 (a,i,I)) = (card I) + 3 by Th12; hence ( m < (card I) + 3 iff m in dom (while<>0 (a,i,I)) ) by AFINSQ_1:66; ::_thesis: verum end; theorem Th14: :: SCPINVAR:10 for a being Int_position for i being Integer for I being Program of SCMPDS holds ( 0 in dom (while<>0 (a,i,I)) & 1 in dom (while<>0 (a,i,I)) ) proof let a be Int_position; ::_thesis: for i being Integer for I being Program of SCMPDS holds ( 0 in dom (while<>0 (a,i,I)) & 1 in dom (while<>0 (a,i,I)) ) let i be Integer; ::_thesis: for I being Program of SCMPDS holds ( 0 in dom (while<>0 (a,i,I)) & 1 in dom (while<>0 (a,i,I)) ) let I be Program of SCMPDS; ::_thesis: ( 0 in dom (while<>0 (a,i,I)) & 1 in dom (while<>0 (a,i,I)) ) 3 <= (card I) + 3 by NAT_1:11; then ( 0 < (card I) + 3 & 1 < (card I) + 3 ) by XXREAL_0:2; hence ( 0 in dom (while<>0 (a,i,I)) & 1 in dom (while<>0 (a,i,I)) ) by Th13; ::_thesis: verum end; theorem Th15: :: SCPINVAR:11 for a being Int_position for i being Integer for I being Program of SCMPDS holds ( (while<>0 (a,i,I)) . 0 = (a,i) <>0_goto 2 & (while<>0 (a,i,I)) . 1 = goto ((card I) + 2) & (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) ) proof let a be Int_position; ::_thesis: for i being Integer for I being Program of SCMPDS holds ( (while<>0 (a,i,I)) . 0 = (a,i) <>0_goto 2 & (while<>0 (a,i,I)) . 1 = goto ((card I) + 2) & (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) ) let i be Integer; ::_thesis: for I being Program of SCMPDS holds ( (while<>0 (a,i,I)) . 0 = (a,i) <>0_goto 2 & (while<>0 (a,i,I)) . 1 = goto ((card I) + 2) & (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) ) let I be Program of SCMPDS; ::_thesis: ( (while<>0 (a,i,I)) . 0 = (a,i) <>0_goto 2 & (while<>0 (a,i,I)) . 1 = goto ((card I) + 2) & (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) ) set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set i3 = goto (- ((card I) + 2)); set I4 = (((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I; set WHL = while<>0 (a,i,I); A1: while<>0 (a,i,I) = (((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' (I ';' (goto (- ((card I) + 2)))) by SCMPDS_4:11; hence (while<>0 (a,i,I)) . 0 = (a,i) <>0_goto 2 by Th5; ::_thesis: ( (while<>0 (a,i,I)) . 1 = goto ((card I) + 2) & (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) ) thus (while<>0 (a,i,I)) . 1 = goto ((card I) + 2) by A1, Th5; ::_thesis: (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) card ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I) = (card (((a,i) <>0_goto 2) ';' (goto ((card I) + 2)))) + (card I) by AFINSQ_1:17 .= (card I) + 2 by SCMP_GCD:5 ; hence (while<>0 (a,i,I)) . ((card I) + 2) = goto (- ((card I) + 2)) by SCMP_GCD:6; ::_thesis: verum end; Lm10: for a being Int_position for i being Integer for I being Program of SCMPDS holds while<>0 (a,i,I) = ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) proof let a be Int_position; ::_thesis: for i being Integer for I being Program of SCMPDS holds while<>0 (a,i,I) = ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) let i be Integer; ::_thesis: for I being Program of SCMPDS holds while<>0 (a,i,I) = ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) let I be Program of SCMPDS; ::_thesis: while<>0 (a,i,I) = ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set i3 = goto (- ((card I) + 2)); thus while<>0 (a,i,I) = (((a,i) <>0_goto 2) ';' ((goto ((card I) + 2)) ';' I)) ';' (goto (- ((card I) + 2))) by SCMPDS_4:16 .= ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) by SCMPDS_4:15 ; ::_thesis: verum end; theorem Th16: :: SCPINVAR:12 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being Program of SCMPDS for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) 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 a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) let s be State of SCMPDS; ::_thesis: for I being Program of SCMPDS for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) let I be Program of SCMPDS; ::_thesis: for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) let a be Int_position; ::_thesis: for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) = 0 implies ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) ) set b = DataLoc ((s . a),i); assume A1: s . (DataLoc ((s . a),i)) = 0 ; ::_thesis: ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) set WHL = while<>0 (a,i,I); set pWH = stop (while<>0 (a,i,I)); set iWH = stop (while<>0 (a,i,I)); set s3 = Initialize s; set P3 = P +* (stop (while<>0 (a,i,I))); set s4 = Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1); set s5 = Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2); set P4 = P +* (stop (while<>0 (a,i,I))); set P5 = P +* (stop (while<>0 (a,i,I))); A3: not DataLoc ((s . a),i) in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; not a in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; then A4: (Initialize s) . (DataLoc (((Initialize s) . a),i)) = (Initialize s) . (DataLoc ((s . a),i)) by FUNCT_4:11 .= 0 by A1, A3, FUNCT_4:11 ; set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set i3 = goto (- ((card I) + 2)); A5: IC (Initialize s) = 0 by MEMSTR_0:47; A6: while<>0 (a,i,I) = ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) by Lm10; Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),(0 + 1)) = Following ((P +* (stop (while<>0 (a,i,I)))),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),0))) by EXTPRO_1:3 .= Following ((P +* (stop (while<>0 (a,i,I)))),(Initialize s)) .= Exec (((a,i) <>0_goto 2),(Initialize s)) by A6, SCMPDS_6:11 ; then A7: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1)) = succ (IC (Initialize s)) by A4, SCMPDS_2:55 .= 0 + 1 by A5 ; A8: stop (while<>0 (a,i,I)) c= P +* (stop (while<>0 (a,i,I))) by FUNCT_4:25; then A9: stop (while<>0 (a,i,I)) c= P +* (stop (while<>0 (a,i,I))) ; A10: 1 in dom (while<>0 (a,i,I)) by Th14; then 1 in dom (stop (while<>0 (a,i,I))) by COMPOS_1:62; then A11: (P +* (stop (while<>0 (a,i,I)))) . 1 = (stop (while<>0 (a,i,I))) . 1 by A9, GRFUNC_1:2 .= (while<>0 (a,i,I)) . 1 by A10, COMPOS_1:63 .= goto ((card I) + 2) by Th15 ; A12: card (while<>0 (a,i,I)) = (card I) + 3 by Th12; then A13: (card I) + 3 in dom (stop (while<>0 (a,i,I))) by COMPOS_1:64; A14: (P +* (stop (while<>0 (a,i,I)))) /. (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) = (P +* (stop (while<>0 (a,i,I)))) . (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) by PBOOLE:143; Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),(1 + 1)) = Following ((P +* (stop (while<>0 (a,i,I)))),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) by EXTPRO_1:3 .= Exec ((goto ((card I) + 2)),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) by A7, A11, A14 ; then A16: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2)) = ICplusConst ((Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1)),((card I) + 2)) by SCMPDS_2:54 .= ((card I) + 2) + 1 by A7, SCMPDS_6:12 .= (card I) + (2 + 1) ; A17: (P +* (stop (while<>0 (a,i,I)))) /. (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2))) = (P +* (stop (while<>0 (a,i,I)))) . (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2))) by PBOOLE:143; stop (while<>0 (a,i,I)) c= P +* (stop (while<>0 (a,i,I))) by A8; then (P +* (stop (while<>0 (a,i,I)))) . ((card I) + 3) = (stop (while<>0 (a,i,I))) . ((card I) + 3) by A13, GRFUNC_1:2 .= halt SCMPDS by A12, COMPOS_1:64 ; then A18: CurInstr ((P +* (stop (while<>0 (a,i,I)))),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2))) = halt SCMPDS by A16, A17; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P_+*_(stop_(while<>0_(a,i,I)))),(Initialize_s),k))_in_dom_(stop_(while<>0_(a,i,I))) let k be Element of NAT ; ::_thesis: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),b1)) in dom (stop (while<>0 (a,i,I))) ( k = 0 or 0 < k ) by NAT_1:3; then A20: ( k = 0 or 0 + 1 <= k ) by INT_1:7; percases ( k = 0 or k = 1 or 1 < k ) by A20, XXREAL_0:1; suppose k = 0 ; ::_thesis: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),b1)) in dom (stop (while<>0 (a,i,I))) then Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),k) = Initialize s by EXTPRO_1:2; hence IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),k)) in dom (stop (while<>0 (a,i,I))) by A5, COMPOS_1:36; ::_thesis: verum end; suppose k = 1 ; ::_thesis: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),b1)) in dom (stop (while<>0 (a,i,I))) hence IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),k)) in dom (stop (while<>0 (a,i,I))) by A10, A7, COMPOS_1:62; ::_thesis: verum end; suppose 1 < k ; ::_thesis: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),b1)) in dom (stop (while<>0 (a,i,I))) then 1 + 1 <= k by INT_1:7; hence IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),k)) in dom (stop (while<>0 (a,i,I))) by A13, A16, A18, EXTPRO_1:5; ::_thesis: verum end; end; end; hence while<>0 (a,i,I) is_closed_on s,P by SCMPDS_6:def_2; ::_thesis: while<>0 (a,i,I) is_halting_on s,P P +* (stop (while<>0 (a,i,I))) halts_on Initialize s by A18, EXTPRO_1:29; hence while<>0 (a,i,I) is_halting_on s,P by SCMPDS_6:def_3; ::_thesis: verum end; theorem Th17: :: SCPINVAR:13 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being Program of SCMPDS for a, c being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being Program of SCMPDS for a, c being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) let s be State of SCMPDS; ::_thesis: for I being Program of SCMPDS for a, c being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) let I be Program of SCMPDS; ::_thesis: for a, c being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) let a, c be Int_position; ::_thesis: for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) = 0 implies IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) ) set b = DataLoc ((s . a),i); assume A1: s . (DataLoc ((s . a),i)) = 0 ; ::_thesis: IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set i3 = goto (- ((card I) + 2)); set WHL = while<>0 (a,i,I); set pWH = stop (while<>0 (a,i,I)); set iWH = stop (while<>0 (a,i,I)); set s3 = Initialize s; set P3 = P +* (stop (while<>0 (a,i,I))); set s4 = Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1); set s5 = Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2); set P4 = P +* (stop (while<>0 (a,i,I))); set P5 = P +* (stop (while<>0 (a,i,I))); A3: IC (Initialize s) = 0 by MEMSTR_0:47; A4: stop (while<>0 (a,i,I)) c= P +* (stop (while<>0 (a,i,I))) by FUNCT_4:25; then A5: stop (while<>0 (a,i,I)) c= P +* (stop (while<>0 (a,i,I))) ; A6: stop (while<>0 (a,i,I)) c= P +* (stop (while<>0 (a,i,I))) by A4; A7: not DataLoc ((s . a),i) in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; A8: while<>0 (a,i,I) = ((a,i) <>0_goto 2) ';' (((goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 2)))) by Lm10; A9: Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),(0 + 1)) = Following ((P +* (stop (while<>0 (a,i,I)))),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),0))) by EXTPRO_1:3 .= Following ((P +* (stop (while<>0 (a,i,I)))),(Initialize s)) .= Exec (((a,i) <>0_goto 2),(Initialize s)) by A8, SCMPDS_6:11 ; A10: 1 in dom (while<>0 (a,i,I)) by Th14; then 1 in dom (stop (while<>0 (a,i,I))) by COMPOS_1:62; then A11: (P +* (stop (while<>0 (a,i,I)))) . 1 = (stop (while<>0 (a,i,I))) . 1 by A5, GRFUNC_1:2 .= (while<>0 (a,i,I)) . 1 by A10, COMPOS_1:63 .= goto ((card I) + 2) by Th15 ; set SAl = Start-At (((card I) + 3),SCMPDS); not a in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; then (Initialize s) . (DataLoc (((Initialize s) . a),i)) = (Initialize s) . (DataLoc ((s . a),i)) by FUNCT_4:11 .= 0 by A1, A7, FUNCT_4:11 ; then A13: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1)) = succ (IC (Initialize s)) by A9, SCMPDS_2:55 .= 0 + 1 by A3 ; A16: (P +* (stop (while<>0 (a,i,I)))) /. (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) = (P +* (stop (while<>0 (a,i,I)))) . (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) by PBOOLE:143; A17: Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),(1 + 1)) = Following ((P +* (stop (while<>0 (a,i,I)))),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) by EXTPRO_1:3 .= Exec ((goto ((card I) + 2)),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1))) by A13, A11, A16 ; then A18: IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2)) = ICplusConst ((Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1)),((card I) + 2)) by SCMPDS_2:54 .= ((card I) + 2) + 1 by A13, SCMPDS_6:12 .= (card I) + (2 + 1) ; A19: (P +* (stop (while<>0 (a,i,I)))) /. (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2))) = (P +* (stop (while<>0 (a,i,I)))) . (IC (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2))) by PBOOLE:143; A20: card (while<>0 (a,i,I)) = (card I) + 3 by Th12; then (card I) + 3 in dom (stop (while<>0 (a,i,I))) by COMPOS_1:64; then (P +* (stop (while<>0 (a,i,I)))) . ((card I) + 3) = (stop (while<>0 (a,i,I))) . ((card I) + 3) by A6, GRFUNC_1:2 .= halt SCMPDS by A20, COMPOS_1:64 ; then A21: CurInstr ((P +* (stop (while<>0 (a,i,I)))),(Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2))) = halt SCMPDS by A18, A19; then P +* (stop (while<>0 (a,i,I))) halts_on Initialize s by EXTPRO_1:29; then A22: Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2) = Result ((P +* (stop (while<>0 (a,i,I)))),(Initialize s)) by A21, EXTPRO_1:def_9; A23: IExec ((while<>0 (a,i,I)),P,(Initialize s)) = Result ((P +* (stop (while<>0 (a,i,I)))),(Initialize s)) by SCMPDS_4:def_5; A24: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((while<>0_(a,i,I)),P,(Initialize_s)))_holds_ (IExec_((while<>0_(a,i,I)),P,(Initialize_s)))_._x_=_(s_+*_(Start-At_(((card_I)_+_3),SCMPDS)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((while<>0 (a,i,I)),P,(Initialize s))) implies (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b1 = (s +* (Start-At (((card I) + 3),SCMPDS))) . b1 ) A25: dom (Start-At (((card I) + 3),SCMPDS)) = {(IC )} by FUNCOP_1:13; assume A26: x in dom (IExec ((while<>0 (a,i,I)),P,(Initialize s))) ; ::_thesis: (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b1 = (s +* (Start-At (((card I) + 3),SCMPDS))) . b1 percases ( x is Int_position or x = IC ) by A26, SCMPDS_4:6; supposeA27: x is Int_position ; ::_thesis: (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b1 = (s +* (Start-At (((card I) + 3),SCMPDS))) . b1 then x <> IC by SCMPDS_2:43; then A28: not x in dom (Start-At (((card I) + 3),SCMPDS)) by A25, TARSKI:def_1; TT: not x in dom (Start-At (0,SCMPDS)) by A27, SCMPDS_4:18; thus (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . x = (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),2)) . x by A22, A23 .= (Comput ((P +* (stop (while<>0 (a,i,I)))),(Initialize s),1)) . x by A17, A27, SCMPDS_2:54 .= (Initialize s) . x by A9, A27, SCMPDS_2:55 .= s . x by TT, FUNCT_4:11 .= (s +* (Start-At (((card I) + 3),SCMPDS))) . x by A28, FUNCT_4:11 ; ::_thesis: verum end; supposeA29: x = IC ; ::_thesis: (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b1 = (s +* (Start-At (((card I) + 3),SCMPDS))) . b1 hence (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . x = (card I) + 3 by A18, A22, A23 .= (s +* (Start-At (((card I) + 3),SCMPDS))) . x by A29, FUNCT_4:113 ; ::_thesis: verum end; end; end; dom (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = the carrier of SCMPDS by PARTFUN1:def_2 .= dom (s +* (Start-At (((card I) + 3),SCMPDS))) by PARTFUN1:def_2 ; hence IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) by A24, FUNCT_1:2; ::_thesis: verum end; theorem :: SCPINVAR:14 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being Program of SCMPDS for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being Program of SCMPDS for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 let s be State of SCMPDS; ::_thesis: for I being Program of SCMPDS for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 let I be Program of SCMPDS; ::_thesis: for a being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 let a be Int_position; ::_thesis: for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) = 0 implies IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 ) assume s . (DataLoc ((s . a),i)) = 0 ; ::_thesis: IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 then IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) by Th17; hence IC (IExec ((while<>0 (a,i,I)),P,(Initialize s))) = (card I) + 3 by FUNCT_4:113; ::_thesis: verum end; theorem Th19: :: SCPINVAR:15 for P being Instruction-Sequence of SCMPDS for s being State of SCMPDS for I being Program of SCMPDS for a, b being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being State of SCMPDS for I being Program of SCMPDS for a, b being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b let s be State of SCMPDS; ::_thesis: for I being Program of SCMPDS for a, b being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b let I be Program of SCMPDS; ::_thesis: for a, b being Int_position for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b let a, b be Int_position; ::_thesis: for i being Integer st s . (DataLoc ((s . a),i)) = 0 holds (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) = 0 implies (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b ) assume s . (DataLoc ((s . a),i)) = 0 ; ::_thesis: (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b then A1: IExec ((while<>0 (a,i,I)),P,(Initialize s)) = s +* (Start-At (((card I) + 3),SCMPDS)) by Th17; not b in dom (Start-At (((card I) + 3),SCMPDS)) by SCMPDS_4:18; hence (IExec ((while<>0 (a,i,I)),P,(Initialize s))) . b = s . b by A1, FUNCT_4:11; ::_thesis: verum end; Lm11: for I being Program of SCMPDS for a being Int_position for i being Integer holds Shift (I,2) c= while<>0 (a,i,I) proof let I be Program of SCMPDS; ::_thesis: for a being Int_position for i being Integer holds Shift (I,2) c= while<>0 (a,i,I) let a be Int_position; ::_thesis: for i being Integer holds Shift (I,2) c= while<>0 (a,i,I) let i be Integer; ::_thesis: Shift (I,2) c= while<>0 (a,i,I) set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set i3 = goto (- ((card I) + 2)); ( card (((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) = 2 & while<>0 (a,i,I) = ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I) ';' (Load (goto (- ((card I) + 2)))) ) by SCMPDS_4:def_3, SCMP_GCD:5; hence Shift (I,2) c= while<>0 (a,i,I) by SCMPDS_7:3; ::_thesis: verum end; registration let I be shiftable Program of SCMPDS; let a be Int_position; let i be Integer; cluster while<>0 (a,i,I) -> shiftable ; correctness coherence while<>0 (a,i,I) is shiftable ; proof set WHL = while<>0 (a,i,I); set i1 = (a,i) <>0_goto 2; set i2 = goto ((card I) + 2); set PF = (((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I; card ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I) = (card (((a,i) <>0_goto 2) ';' (goto ((card I) + 2)))) + (card I) by AFINSQ_1:17 .= 2 + (card I) by SCMP_GCD:5 ; then ( (((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I = ((Load ((a,i) <>0_goto 2)) ';' (Load (goto ((card I) + 2)))) ';' I & (card ((((a,i) <>0_goto 2) ';' (goto ((card I) + 2))) ';' I)) + (- ((card I) + 2)) = 0 ) by SCMPDS_4:def_4; hence while<>0 (a,i,I) is shiftable by SCMPDS_4:23; ::_thesis: verum end; end; registration let I be halt-free Program of SCMPDS; let a be Int_position; let i be Integer; cluster while<>0 (a,i,I) -> halt-free ; correctness coherence while<>0 (a,i,I) is halt-free ; proof reconsider i2 = goto ((card I) + 2) as No-StopCode Instruction of SCMPDS by SCMPDS_5:21; reconsider i3 = goto (- ((card I) + 2)) as No-StopCode Instruction of SCMPDS by SCMPDS_5:21; while<>0 (a,i,I) = ((((a,i) <>0_goto 2) ';' i2) ';' I) ';' i3 ; hence while<>0 (a,i,I) is halt-free ; ::_thesis: verum end; end; begin scheme :: SCPINVAR:sch 3 WhileNHalt{ F1( State of SCMPDS) -> Element of NAT , F2() -> 0 -started State of SCMPDS, F3() -> Instruction-Sequence of SCMPDS, F4() -> halt-free shiftable Program of SCMPDS, F5() -> Int_position, F6() -> Integer, P1[ set ] } : ( while<>0 (F5(),F6(),F4()) is_closed_on F2(),F3() & while<>0 (F5(),F6(),F4()) is_halting_on F2(),F3() ) provided A2: for t being 0 -started State of SCMPDS st P1[t] & F1(t) = 0 holds t . (DataLoc ((F2() . F5()),F6())) = 0 and A3: P1[F2()] and A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = F2() . F5() & t . (DataLoc ((F2() . F5()),F6())) <> 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & F1((Initialize (IExec (F4(),Q,t)))) < F1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) proof I: Initialize F2() = F2() by MEMSTR_0:44; set i1 = (F5(),F6()) <>0_goto 2; set i2 = goto ((card F4()) + 2); set i3 = goto (- ((card F4()) + 2)); set WHL = while<>0 (F5(),F6(),F4()); set pWH = stop (while<>0 (F5(),F6(),F4())); set pI = stop F4(); set b = DataLoc ((F2() . F5()),F6()); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st F1((Initialize t)) <= $1 & P1[ Initialize t] & t . F5() = F2() . F5() holds ( while<>0 (F5(),F6(),F4()) is_closed_on t,Q & while<>0 (F5(),F6(),F4()) is_halting_on t,Q ); A5: 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 A6: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_F1((Initialize_t))_<=_k_+_1_&_P1[_Initialize_t]_&_t_._F5()_=_F2()_._F5()_holds_ (_while<>0_(F5(),F6(),F4())_is_closed_on_t,Q_&_while<>0_(F5(),F6(),F4())_is_halting_on_t,Q_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1((Initialize t)) <= k + 1 & P1[ Initialize t] & t . F5() = F2() . F5() holds ( while<>0 (F5(),F6(),F4()) is_closed_on b2,b3 & while<>0 (F5(),F6(),F4()) is_halting_on b2,b3 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( F1((Initialize t)) <= k + 1 & P1[ Initialize t] & t . F5() = F2() . F5() implies ( while<>0 (F5(),F6(),F4()) is_closed_on b1,b2 & while<>0 (F5(),F6(),F4()) is_halting_on b1,b2 ) ) T: Initialize t = t by MEMSTR_0:44; assume A7: F1((Initialize t)) <= k + 1 ; ::_thesis: ( P1[ Initialize t] & t . F5() = F2() . F5() implies ( while<>0 (F5(),F6(),F4()) is_closed_on b1,b2 & while<>0 (F5(),F6(),F4()) is_halting_on b1,b2 ) ) assume A8: P1[ Initialize t] ; ::_thesis: ( t . F5() = F2() . F5() implies ( while<>0 (F5(),F6(),F4()) is_closed_on b1,b2 & while<>0 (F5(),F6(),F4()) is_halting_on b1,b2 ) ) assume A9: t . F5() = F2() . F5() ; ::_thesis: ( while<>0 (F5(),F6(),F4()) is_closed_on b1,b2 & while<>0 (F5(),F6(),F4()) is_halting_on b1,b2 ) percases ( t . (DataLoc ((F2() . F5()),F6())) = 0 or t . (DataLoc ((F2() . F5()),F6())) <> 0 ) ; suppose t . (DataLoc ((F2() . F5()),F6())) = 0 ; ::_thesis: ( while<>0 (F5(),F6(),F4()) is_closed_on b1,b2 & while<>0 (F5(),F6(),F4()) is_halting_on b1,b2 ) hence ( while<>0 (F5(),F6(),F4()) is_closed_on t,Q & while<>0 (F5(),F6(),F4()) is_halting_on t,Q ) by A9, Th16; ::_thesis: verum end; supposeA10: t . (DataLoc ((F2() . F5()),F6())) <> 0 ; ::_thesis: ( while<>0 (F5(),F6(),F4()) is_closed_on b1,b2 & while<>0 (F5(),F6(),F4()) is_halting_on b1,b2 ) A13: (IExec (F4(),Q,t)) . F5() = t . F5() by A4, A8, A9, A10, T; A14: 0 in dom (stop (while<>0 (F5(),F6(),F4()))) by COMPOS_1:36; A16: while<>0 (F5(),F6(),F4()) = ((F5(),F6()) <>0_goto 2) ';' (((goto ((card F4()) + 2)) ';' F4()) ';' (goto (- ((card F4()) + 2)))) by Lm10; A17: not DataLoc ((F2() . F5()),F6()) in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; set t2 = Initialize t; set Q2 = Q +* (stop F4()); set t3 = Initialize t; set Q3 = Q +* (stop (while<>0 (F5(),F6(),F4()))); set t4 = Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1); set Q4 = Q +* (stop (while<>0 (F5(),F6(),F4()))); A20: stop F4() c= Q +* (stop F4()) by FUNCT_4:25; set m2 = LifeSpan ((Q +* (stop F4())),(Initialize t)); set t5 = Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),(LifeSpan ((Q +* (stop F4())),(Initialize t)))); set Q5 = Q +* (stop (while<>0 (F5(),F6(),F4()))); set l2 = (card F4()) + 2; A21: IC (Initialize t) = 0 by MEMSTR_0:47; set m3 = (LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1; set t6 = Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)); set Q6 = Q +* (stop (while<>0 (F5(),F6(),F4()))); set t7 = Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)); set Q7 = Q +* (stop (while<>0 (F5(),F6(),F4()))); (card F4()) + 2 < (card F4()) + 3 by XREAL_1:6; then A22: (card F4()) + 2 in dom (while<>0 (F5(),F6(),F4())) by Th13; QQ: stop (while<>0 (F5(),F6(),F4())) c= Q +* (stop (while<>0 (F5(),F6(),F4()))) by FUNCT_4:25; while<>0 (F5(),F6(),F4()) c= stop (while<>0 (F5(),F6(),F4())) by AFINSQ_1:74; then A23: while<>0 (F5(),F6(),F4()) c= Q +* (stop (while<>0 (F5(),F6(),F4()))) by QQ, XBOOLE_1:1; Shift (F4(),2) c= while<>0 (F5(),F6(),F4()) by Lm11; then Shift (F4(),2) c= Q +* (stop (while<>0 (F5(),F6(),F4()))) by A23, XBOOLE_1:1; then A24: Shift (F4(),2) c= Q +* (stop (while<>0 (F5(),F6(),F4()))) ; A25: Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(0 + 1)) = Following ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),0))) by EXTPRO_1:3 .= Following ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t)) .= Exec (((F5(),F6()) <>0_goto 2),(Initialize t)) by A16, SCMPDS_6:11 ; for a being Int_position holds (Initialize t) . a = (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)) . a by A25, SCMPDS_2:55; then A27: DataPart (Initialize t) = DataPart (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)) by SCMPDS_4:8; F4() is_halting_on t,Q by A4, A8, A9, A10, T; then A28: Q +* (stop F4()) halts_on Initialize t by SCMPDS_6:def_3; (Q +* (stop F4())) +* (stop F4()) halts_on Initialize (Initialize t) by A28; then A30: F4() is_halting_on Initialize t,Q +* (stop F4()) by SCMPDS_6:def_3; A31: IExec (F4(),Q,(Initialize t)) = Result ((Q +* (stop F4())),(Initialize t)) by SCMPDS_4:def_5; A32: P1[ Initialize (IExec (F4(),Q,(Initialize t)))] by A4, A8, A9, A10, T; A33: F4() is_closed_on t,Q by A4, A8, A9, A10, T; then A34: F4() is_closed_on Initialize t,Q +* (stop F4()) by SCMPDS_6:24; not F5() in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; then (Initialize t) . (DataLoc (((Initialize t) . F5()),F6())) = (Initialize t) . (DataLoc ((F2() . F5()),F6())) by A9, FUNCT_4:11 .= t . (DataLoc ((F2() . F5()),F6())) by A17, FUNCT_4:11 ; then A35: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)) = ICplusConst ((Initialize t),2) by A10, A25, SCMPDS_2:55 .= 0 + 2 by A21, SCMPDS_6:12 ; then A36: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),(LifeSpan ((Q +* (stop F4())),(Initialize t))))) = (card F4()) + 2 by A20, A30, A34, A27, A24, SCMPDS_7:18; A37: (Q +* (stop (while<>0 (F5(),F6(),F4())))) /. (IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)))) = (Q +* (stop (while<>0 (F5(),F6(),F4())))) . (IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)))) by PBOOLE:143; A38: Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)) = Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),(LifeSpan ((Q +* (stop F4())),(Initialize t)))) by EXTPRO_1:4; then A39: CurInstr ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)))) = (Q +* (stop (while<>0 (F5(),F6(),F4())))) . ((card F4()) + 2) by A20, A30, A34, A35, A27, A24, A37, SCMPDS_7:18 .= (Q +* (stop (while<>0 (F5(),F6(),F4())))) . ((card F4()) + 2) .= (Q +* (stop (while<>0 (F5(),F6(),F4())))) . ((card F4()) + 2) .= (while<>0 (F5(),F6(),F4())) . ((card F4()) + 2) by A22, A23, GRFUNC_1:2 .= goto (- ((card F4()) + 2)) by Th15 ; A41: Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)) = Following ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card F4()) + 2))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1)))) by A39 ; then IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1))),(0 - ((card F4()) + 2))) by SCMPDS_2:54 .= 0 by A36, A38, SCMPDS_7:1 ; then B42: Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))) = Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)) by MEMSTR_0:46; A42: (Q +* (stop (while<>0 (F5(),F6(),F4())))) +* (stop (while<>0 (F5(),F6(),F4()))) = Q +* (stop (while<>0 (F5(),F6(),F4()))) ; A43: DataPart (Comput ((Q +* (stop F4())),(Initialize t),(LifeSpan ((Q +* (stop F4())),(Initialize t))))) = DataPart (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),(LifeSpan ((Q +* (stop F4())),(Initialize t))))) by A20, A30, A34, A35, A27, A24, SCMPDS_7:18; then A44: DataPart (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),(LifeSpan ((Q +* (stop F4())),(Initialize t))))) = DataPart (Result ((Q +* (stop F4())),(Initialize t))) by A28, EXTPRO_1:23 .= DataPart (Result ((Q +* (stop F4())),(Initialize t))) .= DataPart (IExec (F4(),Q,(Initialize t))) by SCMPDS_4:def_5 ; InsCode (goto (- ((card F4()) + 2))) = 14 by SCMPDS_2:12; then InsCode (goto (- ((card F4()) + 2))) in {0,4,5,6,14} by ENUMSET1:def_3; then A45: Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))) = Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1))) by A41, SCMPDS_8:3 .= Initialize (IExec (F4(),Q,(Initialize t))) by A44, A38, MEMSTR_0:80 ; A46: now__::_thesis:_not_F1((Initialize_(Comput_((Q_+*_(stop_(while<>0_(F5(),F6(),F4())))),(Initialize_t),(((LifeSpan_((Q_+*_(stop_F4())),(Initialize_t)))_+_1)_+_1)))))_>_k F1((Initialize (IExec (F4(),Q,(Initialize t))))) < F1((Initialize t)) by A4, A8, A9, A10, T; then A47: F1((Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))))) < k + 1 by A7, A45, XXREAL_0:2; assume F1((Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))))) > k ; ::_thesis: contradiction hence contradiction by A47, INT_1:7; ::_thesis: verum end; A48: (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),(LifeSpan ((Q +* (stop F4())),(Initialize t))))) . F5() = (Comput ((Q +* (stop F4())),(Initialize t),(LifeSpan ((Q +* (stop F4())),(Initialize t))))) . F5() by A43, SCMPDS_4:8 .= (Result ((Q +* (stop F4())),(Initialize t))) . F5() by A28, EXTPRO_1:23 .= F2() . F5() by A9, A13, A31, T ; A50: (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))) . F5() = (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1))) . F5() by A41, SCMPDS_2:54 .= F2() . F5() by A48, EXTPRO_1:4 ; X1: (Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)))) . F5() = (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))) . F5() by SCMPDS_5:15; P1[ Initialize (Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))))] by A32, A45; then X3: ( while<>0 (F5(),F6(),F4()) is_closed_on Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))),Q +* (stop (while<>0 (F5(),F6(),F4()))) & while<>0 (F5(),F6(),F4()) is_halting_on Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))),Q +* (stop (while<>0 (F5(),F6(),F4()))) ) by A6, A46, A50, X1; A51: while<>0 (F5(),F6(),F4()) is_closed_on Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)),Q +* (stop (while<>0 (F5(),F6(),F4()))) proof for k being Element of NAT holds IC (Comput (((Q +* (stop (while<>0 (F5(),F6(),F4())))) +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1))))),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) by X3, SCMPDS_6:def_2; hence while<>0 (F5(),F6(),F4()) is_closed_on Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)),Q +* (stop (while<>0 (F5(),F6(),F4()))) by SCMPDS_6:def_2; ::_thesis: verum end; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(while<>0_(F5(),F6(),F4())))),(Initialize_t),k))_in_dom_(stop_(while<>0_(F5(),F6(),F4()))) let k be Element of NAT ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),b1)) in dom (stop (while<>0 (F5(),F6(),F4()))) percases ( k < ((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),b1)) in dom (stop (while<>0 (F5(),F6(),F4()))) then A52: k <= (LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1 by INT_1:7; hereby ::_thesis: verum percases ( k <= LifeSpan ((Q +* (stop F4())),(Initialize t)) or k = (LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1 ) by A52, NAT_1:8; supposeA53: k <= LifeSpan ((Q +* (stop F4())),(Initialize t)) ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) hereby ::_thesis: verum percases ( k = 0 or k <> 0 ) ; suppose k = 0 ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) hence IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) by A14, A21, EXTPRO_1:2; ::_thesis: verum end; suppose k <> 0 ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) then consider kn being Nat such that A54: k = kn + 1 by NAT_1:6; reconsider kn = kn as Element of NAT by ORDINAL1:def_12; reconsider lm = IC (Comput ((Q +* (stop F4())),(Initialize t),kn)) as Element of NAT ; kn < k by A54, XREAL_1:29; then kn < LifeSpan ((Q +* (stop F4())),(Initialize t)) by A53, XXREAL_0:2; then (IC (Comput ((Q +* (stop F4())),(Initialize t),kn))) + 2 = IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),1)),kn)) by A20, A30, A34, A35, A27, A24, SCMPDS_7:16; then A56: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) = lm + 2 by A54, EXTPRO_1:4; IC (Comput ((Q +* (stop F4())),(Initialize t),kn)) in dom (stop F4()) by A33, SCMPDS_6:def_2; then lm < card (stop F4()) by AFINSQ_1:66; then lm < (card F4()) + 1 by COMPOS_1:55; then A57: lm + 2 < ((card F4()) + 1) + 2 by XREAL_1:6; (card F4()) + 3 < (card F4()) + 4 by XREAL_1:6; then lm + 2 < (card F4()) + 4 by A57, XXREAL_0:2; then lm + 2 < card (stop (while<>0 (F5(),F6(),F4()))) by Lm9; hence IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) by A56, AFINSQ_1:66; ::_thesis: verum end; end; end; end; supposeA58: k = (LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) (card F4()) + 2 in dom (stop (while<>0 (F5(),F6(),F4()))) by A22, COMPOS_1:62; hence IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) by A20, A30, A34, A35, A27, A24, A38, A58, SCMPDS_7:18; ::_thesis: verum end; end; end; end; suppose k >= ((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),b1)) in dom (stop (while<>0 (F5(),F6(),F4()))) then consider nn being Nat such that A59: k = (((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k) = Comput (((Q +* (stop (while<>0 (F5(),F6(),F4())))) +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)))),nn) by A59, B42, EXTPRO_1:4; hence IC (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),k)) in dom (stop (while<>0 (F5(),F6(),F4()))) by A51, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence while<>0 (F5(),F6(),F4()) is_closed_on t,Q by SCMPDS_6:def_2; ::_thesis: while<>0 (F5(),F6(),F4()) is_halting_on t,Q while<>0 (F5(),F6(),F4()) is_halting_on Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)),Q +* (stop (while<>0 (F5(),F6(),F4()))) proof (Q +* (stop (while<>0 (F5(),F6(),F4())))) +* (stop (while<>0 (F5(),F6(),F4()))) halts_on Initialize (Initialize (Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)))) by X3, SCMPDS_6:def_3; hence while<>0 (F5(),F6(),F4()) is_halting_on Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)),Q +* (stop (while<>0 (F5(),F6(),F4()))) by SCMPDS_6:def_3; ::_thesis: verum end; then Q +* (stop (while<>0 (F5(),F6(),F4()))) halts_on Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)) by A42, B42, SCMPDS_6:def_3; then Q +* (stop (while<>0 (F5(),F6(),F4()))) halts_on Comput ((Q +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize t),(((LifeSpan ((Q +* (stop F4())),(Initialize t))) + 1) + 1)) ; then Q +* (stop (while<>0 (F5(),F6(),F4()))) halts_on Initialize t by EXTPRO_1:22; hence while<>0 (F5(),F6(),F4()) is_halting_on t,Q by SCMPDS_6:def_3; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; set n = F1(F2()); A60: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1((Initialize t)) <= 0 & P1[ Initialize t] & t . F5() = F2() . F5() holds ( while<>0 (F5(),F6(),F4()) is_closed_on t,Q & while<>0 (F5(),F6(),F4()) is_halting_on t,Q ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( F1((Initialize t)) <= 0 & P1[ Initialize t] & t . F5() = F2() . F5() implies ( while<>0 (F5(),F6(),F4()) is_closed_on t,Q & while<>0 (F5(),F6(),F4()) is_halting_on t,Q ) ) assume that A61: ( F1((Initialize t)) <= 0 & P1[ Initialize t] ) and A62: t . F5() = F2() . F5() ; ::_thesis: ( while<>0 (F5(),F6(),F4()) is_closed_on t,Q & while<>0 (F5(),F6(),F4()) is_halting_on t,Q ) X: F1((Initialize t)) >= 0 by NAT_1:2; (Initialize t) . (DataLoc ((F2() . F5()),F6())) = t . (DataLoc ((F2() . F5()),F6())) by SCMPDS_5:15; then t . (DataLoc ((F2() . F5()),F6())) = 0 by A2, A61, X, XXREAL_0:1; hence ( while<>0 (F5(),F6(),F4()) is_closed_on t,Q & while<>0 (F5(),F6(),F4()) is_halting_on t,Q ) by A62, Th16; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A60, A5); then S1[F1(F2())] ; hence ( while<>0 (F5(),F6(),F4()) is_closed_on F2(),F3() & while<>0 (F5(),F6(),F4()) is_halting_on F2(),F3() ) by A3, I; ::_thesis: verum end; scheme :: SCPINVAR:sch 4 WhileNExec{ F1( State of SCMPDS) -> Element of NAT , F2() -> 0 -started State of SCMPDS, F3() -> Instruction-Sequence of SCMPDS, F4() -> halt-free shiftable Program of SCMPDS, F5() -> Int_position, F6() -> Integer, P1[ set ] } : IExec ((while<>0 (F5(),F6(),F4())),F3(),F2()) = IExec ((while<>0 (F5(),F6(),F4())),F3(),(Initialize (IExec (F4(),F3(),F2())))) provided A2: F2() . (DataLoc ((F2() . F5()),F6())) <> 0 and A3: for t being 0 -started State of SCMPDS st P1[t] & F1(t) = 0 holds t . (DataLoc ((F2() . F5()),F6())) = 0 and A4: P1[F2()] and A5: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = F2() . F5() & t . (DataLoc ((F2() . F5()),F6())) <> 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & F1((Initialize (IExec (F4(),Q,t)))) < F1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) proof I: Initialize F2() = F2() by MEMSTR_0:44; set WHL = while<>0 (F5(),F6(),F4()); set s1 = F2(); set P1 = F3() +* (stop (while<>0 (F5(),F6(),F4()))); set sI = F2(); set PI = F3() +* (stop F4()); set m1 = (LifeSpan ((F3() +* (stop F4())),F2())) + 2; set s2 = Initialize (IExec (F4(),F3(),F2())); set m2 = LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))); A9: stop F4() c= F3() +* (stop F4()) by FUNCT_4:25; F4() is_closed_on F2(),F3() by A2, A4, A5; then A10: F4() is_closed_on F2(),F3() +* (stop F4()) by I, SCMPDS_6:24; F4() is_halting_on F2(),F3() by A2, A4, A5; then A11: F3() +* (stop F4()) halts_on F2() by I, SCMPDS_6:def_3; (F3() +* (stop F4())) +* (stop F4()) halts_on Initialize F2() by A11, I; then A13: F4() is_halting_on F2(),F3() +* (stop F4()) by SCMPDS_6:def_3; set s4 = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1); set P4 = F3() +* (stop (while<>0 (F5(),F6(),F4()))); set i1 = (F5(),F6()) <>0_goto 2; set i2 = goto ((card F4()) + 2); set i3 = goto (- ((card F4()) + 2)); set b = DataLoc ((F2() . F5()),F6()); A14: IC F2() = 0 by I, MEMSTR_0:47; set mI = LifeSpan ((F3() +* (stop F4())),F2()); set s5 = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)),(LifeSpan ((F3() +* (stop F4())),F2()))); set P5 = F3() +* (stop (while<>0 (F5(),F6(),F4()))); set l2 = (card F4()) + 2; A15: while<>0 (F5(),F6(),F4()) = ((F5(),F6()) <>0_goto 2) ';' (((goto ((card F4()) + 2)) ';' F4()) ';' (goto (- ((card F4()) + 2)))) by Lm10; set m3 = (LifeSpan ((F3() +* (stop F4())),F2())) + 1; set s6 = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)); set P6 = F3() +* (stop (while<>0 (F5(),F6(),F4()))); (card F4()) + 2 < (card F4()) + 3 by XREAL_1:6; then A16: (card F4()) + 2 in dom (while<>0 (F5(),F6(),F4())) by Th13; set s7 = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1)); set P7 = F3() +* (stop (while<>0 (F5(),F6(),F4()))); A17: IExec (F4(),F3(),F2()) = Result ((F3() +* (stop F4())),F2()) by SCMPDS_4:def_5; QQ: stop (while<>0 (F5(),F6(),F4())) c= F3() +* (stop (while<>0 (F5(),F6(),F4()))) by FUNCT_4:25; while<>0 (F5(),F6(),F4()) c= stop (while<>0 (F5(),F6(),F4())) by AFINSQ_1:74; then A19: while<>0 (F5(),F6(),F4()) c= F3() +* (stop (while<>0 (F5(),F6(),F4()))) by QQ, XBOOLE_1:1; deffunc H1( State of SCMPDS) -> Element of NAT = F1($1); A20: for t being 0 -started State of SCMPDS st P1[t] & H1(t) = 0 holds t . (DataLoc ((F2() . F5()),F6())) = 0 by A3; A21: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = F2() . F5() & t . (DataLoc ((F2() . F5()),F6())) <> 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & H1( Initialize (IExec (F4(),Q,t))) < H1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) by A5; A22: Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(0 + 1)) = Following ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),0))) by EXTPRO_1:3 .= Following ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()) .= Exec (((F5(),F6()) <>0_goto 2),F2()) by A15, I, SCMPDS_6:11 ; Shift (F4(),2) c= while<>0 (F5(),F6(),F4()) by Lm11; then Shift (F4(),2) c= F3() +* (stop (while<>0 (F5(),F6(),F4()))) by A19, XBOOLE_1:1; then A23: Shift (F4(),2) c= F3() +* (stop (while<>0 (F5(),F6(),F4()))) ; for a being Int_position holds F2() . a = (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)) . a by A22, SCMPDS_2:55; then A25: DataPart F2() = DataPart (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)) by SCMPDS_4:8; A26: IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)) = ICplusConst (F2(),2) by A2, A22, SCMPDS_2:55 .= 0 + 2 by A14, SCMPDS_6:12 ; then A27: IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)),(LifeSpan ((F3() +* (stop F4())),F2())))) = (card F4()) + 2 by A9, A13, A10, A25, A23, SCMPDS_7:18; A28: (F3() +* (stop (while<>0 (F5(),F6(),F4())))) /. (IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) = (F3() +* (stop (while<>0 (F5(),F6(),F4())))) . (IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) by PBOOLE:143; A29: Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)) = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)),(LifeSpan ((F3() +* (stop F4())),F2()))) by EXTPRO_1:4; then A30: CurInstr ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) = (F3() +* (stop (while<>0 (F5(),F6(),F4())))) . ((card F4()) + 2) by A9, A13, A10, A26, A25, A23, A28, SCMPDS_7:18 .= (F3() +* (stop (while<>0 (F5(),F6(),F4())))) . ((card F4()) + 2) .= (F3() +* (stop (while<>0 (F5(),F6(),F4())))) . ((card F4()) + 2) .= (while<>0 (F5(),F6(),F4())) . ((card F4()) + 2) by A16, A19, GRFUNC_1:2 .= goto (- ((card F4()) + 2)) by Th15 ; A32: Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1)) = Following ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card F4()) + 2))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) by A30 ; then IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1))) = ICplusConst ((Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1))),(0 - ((card F4()) + 2))) by SCMPDS_2:54 .= 0 by A27, A29, SCMPDS_7:1 ; then A33: IC (Initialize (IExec (F4(),F3(),F2()))) = IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 2))) by MEMSTR_0:47; A34: DataPart (Comput ((F3() +* (stop F4())),F2(),(LifeSpan ((F3() +* (stop F4())),F2())))) = DataPart (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)),(LifeSpan ((F3() +* (stop F4())),F2())))) by A9, A13, A10, A26, A25, A23, SCMPDS_7:18; now__::_thesis:_for_x_being_Int_position_holds_(Comput_((F3()_+*_(stop_(while<>0_(F5(),F6(),F4())))),F2(),(((LifeSpan_((F3()_+*_(stop_F4())),F2()))_+_1)_+_1)))_._x_=_(Initialize_(IExec_(F4(),F3(),F2())))_._x let x be Int_position; ::_thesis: (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1))) . x = (Initialize (IExec (F4(),F3(),F2()))) . x A35: not x in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),1)),(LifeSpan ((F3() +* (stop F4())),F2())))) . x = (Comput ((F3() +* (stop F4())),F2(),(LifeSpan ((F3() +* (stop F4())),F2())))) . x by A34, SCMPDS_4:8 .= (Result ((F3() +* (stop F4())),F2())) . x by A11, EXTPRO_1:23 .= (IExec (F4(),F3(),F2())) . x by A17 ; hence (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1))) . x = (IExec (F4(),F3(),F2())) . x by A29, A32, SCMPDS_2:54 .= (Initialize (IExec (F4(),F3(),F2()))) . x by A35, FUNCT_4:11 ; ::_thesis: verum end; then A37: DataPart (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1))) = DataPart (Initialize (IExec (F4(),F3(),F2()))) by SCMPDS_4:8; set m0 = LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()); A40: P1[F2()] by A4; ( while<>0 (F5(),F6(),F4()) is_closed_on F2(),F3() & while<>0 (F5(),F6(),F4()) is_halting_on F2(),F3() ) from SCPINVAR:sch_3(A20, A40, A21); then A41: F3() +* (stop (while<>0 (F5(),F6(),F4()))) halts_on F2() by I, SCMPDS_6:def_3; deffunc H2( State of SCMPDS) -> Element of NAT = H1($1); set Es = IExec (F4(),F3(),F2()); set bj = DataLoc (((IExec (F4(),F3(),F2())) . F5()),F6()); set EP = F3(); A42: (IExec (F4(),F3(),F2())) . F5() = F2() . F5() by A2, A4, A5; U1: (Initialize (IExec (F4(),F3(),F2()))) . F5() = (IExec (F4(),F3(),F2())) . F5() by SCMPDS_5:15; A43: for t being 0 -started State of SCMPDS st P1[t] & H2(t) = 0 holds t . (DataLoc (((Initialize (IExec (F4(),F3(),F2()))) . F5()),F6())) = 0 by A3, A42, U1; A45: P1[ Initialize (IExec (F4(),F3(),F2()))] by A2, A4, A5; A44: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = (Initialize (IExec (F4(),F3(),F2()))) . F5() & t . (DataLoc (((Initialize (IExec (F4(),F3(),F2()))) . F5()),F6())) <> 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & H2( Initialize (IExec (F4(),Q,t))) < H2(t) & P1[ Initialize (IExec (F4(),Q,t))] ) by A5, A42, U1; XXX: ( while<>0 (F5(),F6(),F4()) is_closed_on Initialize (IExec (F4(),F3(),F2())),F3() & while<>0 (F5(),F6(),F4()) is_halting_on Initialize (IExec (F4(),F3(),F2())),F3() ) from SCPINVAR:sch_3(A43, A45, A44); while<>0 (F5(),F6(),F4()) is_halting_on IExec (F4(),F3(),F2()),F3() proof F3() +* (stop (while<>0 (F5(),F6(),F4()))) halts_on Initialize (Initialize (IExec (F4(),F3(),F2()))) by XXX, SCMPDS_6:def_3; hence while<>0 (F5(),F6(),F4()) is_halting_on IExec (F4(),F3(),F2()),F3() by SCMPDS_6:def_3; ::_thesis: verum end; then A46: F3() +* (stop (while<>0 (F5(),F6(),F4()))) halts_on Initialize (IExec (F4(),F3(),F2())) by SCMPDS_6:def_3; A48: Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 2)) = Initialize (IExec (F4(),F3(),F2())) by A37, A33, MEMSTR_0:78; then CurInstr ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 2)))) = (F5(),F6()) <>0_goto 2 by A15, SCMPDS_6:11; then LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()) > (LifeSpan ((F3() +* (stop F4())),F2())) + 2 by A41, EXTPRO_1:36, SCMPDS_6:16; then consider nn being Nat such that A49: LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()) = ((LifeSpan ((F3() +* (stop F4())),F2())) + 2) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 2) + (LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))))))) = IC (Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))),(LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))))))) by A48, EXTPRO_1:4; then CurInstr ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 2) + (LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))))))))) = CurInstr ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))),(LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))))))) .= halt SCMPDS by A46, EXTPRO_1:def_15 ; then ((LifeSpan ((F3() +* (stop F4())),F2())) + 2) + (LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))))) >= LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()) by A41, EXTPRO_1:def_15; then A51: LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))) >= nn by A49, XREAL_1:6; A52: Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2(),(LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()))) = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))),nn) by A48, A49, EXTPRO_1:4; then CurInstr ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))),nn))) = halt SCMPDS by A41, EXTPRO_1:def_15; then nn >= LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))) by A46, EXTPRO_1:def_15; then nn = LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))) by A51, XXREAL_0:1; then Result ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),F2()) = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))),(LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))))) by A41, A52, EXTPRO_1:23; hence IExec ((while<>0 (F5(),F6(),F4())),F3(),F2()) = Comput ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2()))),(LifeSpan ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))))) by SCMPDS_4:def_5 .= Result ((F3() +* (stop (while<>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))) by A46, EXTPRO_1:23 .= IExec ((while<>0 (F5(),F6(),F4())),F3(),(Initialize (IExec (F4(),F3(),F2())))) by SCMPDS_4:def_5 ; ::_thesis: verum end; scheme :: SCPINVAR:sch 5 WhileNEnd{ F1( State of SCMPDS) -> Element of NAT , F2() -> 0 -started State of SCMPDS, F3() -> Instruction-Sequence of SCMPDS, F4() -> halt-free shiftable Program of SCMPDS, F5() -> Int_position, F6() -> Integer, P1[ set ] } : ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),F3(),F2())))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),F3(),F2()))] ) provided A2: for t being 0 -started State of SCMPDS st P1[t] holds ( F1(t) = 0 iff t . (DataLoc ((F2() . F5()),F6())) = 0 ) and A3: P1[F2()] and A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = F2() . F5() & t . (DataLoc ((F2() . F5()),F6())) <> 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & F1((Initialize (IExec (F4(),Q,t)))) < F1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) proof set b = DataLoc ((F2() . F5()),F6()); set WHL = while<>0 (F5(),F6(),F4()); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st F1(t) <= $1 & t . F5() = F2() . F5() & P1[t] holds ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t))] ); A5: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1(t) <= 0 & t . F5() = F2() . F5() & P1[t] holds ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t))] ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(t) <= 0 & t . F5() = F2() . F5() & P1[t] implies ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t))] ) ) T: Initialize t = t by MEMSTR_0:44; assume that A6: F1(t) <= 0 and A7: t . F5() = F2() . F5() and A8: P1[t] ; ::_thesis: ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t))] ) F1(t) >= 0 by NAT_1:2; then A9: F1(t) = 0 by A6, XXREAL_0:1; then t . (DataLoc ((t . F5()),F6())) = 0 by A2, A7, A8; then for b being Int_position holds (IExec ((while<>0 (F5(),F6(),F4())),Q,t)) . b = t . b by Th19, T; hence ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),Q,t))] ) by A8, A9, T, SCPISORT:4; ::_thesis: verum end; A10: 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 A11: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_u_being_0_-started_State_of_SCMPDS for_U_being_Instruction-Sequence_of_SCMPDS_st_F1(u)_<=_k_+_1_&_u_._F5()_=_F2()_._F5()_&_P1[u]_holds_ (_F1((Initialize_(IExec_((while<>0_(F5(),F6(),F4())),U,u))))_=_0_&_P1[_Initialize_(IExec_((while<>0_(F5(),F6(),F4())),U,u))]_) let u be 0 -started State of SCMPDS; ::_thesis: for U being Instruction-Sequence of SCMPDS st F1(u) <= k + 1 & u . F5() = F2() . F5() & P1[u] holds ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),b3,b2)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),b3,b2))] ) let U be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(u) <= k + 1 & u . F5() = F2() . F5() & P1[u] implies ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1))] ) ) assume that A12: F1(u) <= k + 1 and A13: u . F5() = F2() . F5() and A14: P1[u] ; ::_thesis: ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1))] ) percases ( F1(u) = 0 or F1(u) <> 0 ) ; suppose F1(u) = 0 ; ::_thesis: ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1))] ) hence ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),U,u)))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),U,u))] ) by A5, A13, A14; ::_thesis: verum end; supposeA15: F1(u) <> 0 ; ::_thesis: ( H1( Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),b2,b1))] ) set Iu = IExec (F4(),U,u); A16: u . (DataLoc ((F2() . F5()),F6())) <> 0 by A2, A14, A15; then A17: ( (IExec (F4(),U,u)) . F5() = F2() . F5() & P1[ Initialize (IExec (F4(),U,u))] ) by A4, A13, A14; deffunc H1( State of SCMPDS) -> Element of NAT = F1($1); A18: P1[u] by A14; A19: for t being 0 -started State of SCMPDS st P1[t] & H1(t) = 0 holds t . (DataLoc ((u . F5()),F6())) = 0 by A2, A13; H1( Initialize (IExec (F4(),U,u))) < H1(u) by A4, A13, A14, A16; then H1( Initialize (IExec (F4(),U,u))) + 1 <= H1(u) by INT_1:7; then H1( Initialize (IExec (F4(),U,u))) + 1 <= k + 1 by A12, XXREAL_0:2; then A20: H1( Initialize (IExec (F4(),U,u))) <= k by XREAL_1:6; A21: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st P1[t] & t . F5() = u . F5() & t . (DataLoc ((u . F5()),F6())) <> 0 holds ( (IExec (F4(),Q,t)) . F5() = t . F5() & F4() is_closed_on t,Q & F4() is_halting_on t,Q & H1( Initialize (IExec (F4(),Q,t))) < H1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) by A4, A13; A22: u . (DataLoc ((u . F5()),F6())) <> 0 by A2, A13, A14, A15; XX: IExec ((while<>0 (F5(),F6(),F4())),U,u) = IExec ((while<>0 (F5(),F6(),F4())),U,(Initialize (IExec (F4(),U,u)))) from SCPINVAR:sch_4(A22, A19, A18, A21); (Initialize (IExec (F4(),U,u))) . F5() = (IExec (F4(),U,u)) . F5() by SCMPDS_5:15; hence ( H1( Initialize (IExec ((while<>0 (F5(),F6(),F4())),U,u))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),U,u))] ) by A11, A20, A17, XX; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A5, A10); then S1[F1(F2())] ; hence ( F1((Initialize (IExec ((while<>0 (F5(),F6(),F4())),F3(),F2())))) = 0 & P1[ Initialize (IExec ((while<>0 (F5(),F6(),F4())),F3(),F2()))] ) by A3; ::_thesis: verum end; theorem Th20: :: SCPINVAR:16 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, b, c being Int_position for i, d being Integer st card I > 0 & s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,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, b, c being Int_position for i, d being Integer st card I > 0 & s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free shiftable Program of SCMPDS for a, b, c being Int_position for i, d being Integer st card I > 0 & s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a, b, c being Int_position for i, d being Integer st card I > 0 & s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) let a, b, c be Int_position; ::_thesis: for i, d being Integer st card I > 0 & s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) let i, d be Integer; ::_thesis: ( card I > 0 & s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) implies ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) ) set ci = DataLoc ((s . a),i); assume card I > 0 ; ::_thesis: ( not s . a = d or not s . b > 0 or not s . c > 0 or not s . (DataLoc (d,i)) = (s . b) - (s . c) or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c & not ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) or ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) ) consider f being Function of (product (the_Values_of SCMPDS)),NAT such that A2: for s being State of SCMPDS holds ( ( s . b = s . c implies f . s = 0 ) & ( s . b <> s . c implies f . s = max ((abs (s . b)),(abs (s . c))) ) ) by Th6; deffunc H1( State of SCMPDS) -> Element of NAT = f . $1; defpred S1[ set ] means ex t being State of SCMPDS st ( t = $1 & t . b > 0 & t . c > 0 & t . (DataLoc (d,i)) = (t . b) - (t . c) ); assume that A3: s . a = d and A4: s . b > 0 and A5: s . c > 0 and A6: s . (DataLoc (d,i)) = (s . b) - (s . c) ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c & not ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) or ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) ) assume A7: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ; ::_thesis: ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P & ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) ) A8: 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,Q,t))_._a_=_t_._a_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_H1(_Initialize_(IExec_(I,Q,t)))_<_H1(t)_&_S1[_Initialize_(IExec_(I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) ) assume that A9: S1[t] and A10: t . a = s . a and A11: t . (DataLoc ((s . a),i)) <> 0 ; ::_thesis: ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) set It = IExec (I,Q,t); set t2 = Initialize (IExec (I,Q,t)); set t1 = t; set x = (IExec (I,Q,t)) . b; set y = (IExec (I,Q,t)) . c; consider v being State of SCMPDS such that A12: v = t and A13: v . b > 0 and A14: v . c > 0 and A15: v . (DataLoc (d,i)) = (v . b) - (v . c) by A9; A16: t . b > 0 by A12, A13; A17: t . c > 0 by A12, A14; A18: t . (DataLoc (d,i)) = (v . b) - (v . c) by A12, A15 .= (t . b) - (v . c) by A12 .= (t . b) - (t . c) by A12 ; then A19: t . b <> t . c by A3, A11; hence (IExec (I,Q,t)) . a = t . a by A3, A7, A10, A16, A17, A18; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) thus ( I is_closed_on t,Q & I is_halting_on t,Q ) by A3, A7, A10, A16, A17, A18, A19; ::_thesis: ( H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) A20: now__::_thesis:_(_(IExec_(I,Q,t))_._b_>_0_&_(IExec_(I,Q,t))_._c_>_0_&_max_(((IExec_(I,Q,t))_._b),((IExec_(I,Q,t))_._c))_<_max_((t_._b),(t_._c))_) percases ( t . b > t . c or t . b <= t . c ) ; supposeA21: t . b > t . c ; ::_thesis: ( (IExec (I,Q,t)) . b > 0 & (IExec (I,Q,t)) . c > 0 & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) then (t . b) - (t . c) > 0 by XREAL_1:50; hence (IExec (I,Q,t)) . b > 0 by A3, A7, A10, A16, A17, A18, A21; ::_thesis: ( (IExec (I,Q,t)) . c > 0 & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) thus (IExec (I,Q,t)) . c > 0 by A3, A7, A10, A16, A17, A18, A21; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) A22: (IExec (I,Q,t)) . b = (t . b) - (t . c) by A3, A7, A10, A16, A17, A18, A21; hereby ::_thesis: verum A23: max ((t . b),(t . c)) = t . b by A21, XXREAL_0:def_10; percases ( max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b or max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c ) by XXREAL_0:16; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A17, A22, A23, XREAL_1:44; ::_thesis: verum end; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A3, A7, A10, A16, A17, A18, A21, A23; ::_thesis: verum end; end; end; end; supposeA24: t . b <= t . c ; ::_thesis: ( (IExec (I,Q,t)) . b > 0 & (IExec (I,Q,t)) . c > 0 & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) hence (IExec (I,Q,t)) . b > 0 by A3, A7, A10, A16, A17, A18, A19; ::_thesis: ( (IExec (I,Q,t)) . c > 0 & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) t . b < t . c by A19, A24, XXREAL_0:1; then (t . c) - (t . b) > 0 by XREAL_1:50; hence (IExec (I,Q,t)) . c > 0 by A3, A7, A10, A16, A17, A18, A19, A24; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) A25: (IExec (I,Q,t)) . c = (t . c) - (t . b) by A3, A7, A10, A16, A17, A18, A19, A24; A26: (IExec (I,Q,t)) . b = t . b by A3, A7, A10, A16, A17, A18, A19, A24; hereby ::_thesis: verum A27: max ((t . b),(t . c)) = t . c by A24, XXREAL_0:def_10; percases ( max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c or max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b ) by XXREAL_0:16; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A16, A25, A27, XREAL_1:44; ::_thesis: verum end; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A19, A24, A26, A27, XXREAL_0:1; ::_thesis: verum end; end; end; end; end; end; thus H1( Initialize (IExec (I,Q,t))) < H1(t) ::_thesis: S1[ Initialize (IExec (I,Q,t))] proof t . b <> t . c by A19; then t . b <> t . c ; then A28: H1(t) = max ((abs (t . b)),(abs (t . c))) by A2 .= max ((abs (t . b)),(abs (t . c))) .= max ((abs (t . b)),(abs (t . c))) .= max ((t . b),(abs (t . c))) by A16, ABSVALUE:def_1 .= max ((t . b),(t . c)) by A17, ABSVALUE:def_1 ; then H1(t) >= t . b by XXREAL_0:25; then A29: H1(t) > 0 by A16, XXREAL_0:2; percases ( (Initialize (IExec (I,Q,t))) . b = (Initialize (IExec (I,Q,t))) . c or (Initialize (IExec (I,Q,t))) . b <> (Initialize (IExec (I,Q,t))) . c ) ; suppose (Initialize (IExec (I,Q,t))) . b = (Initialize (IExec (I,Q,t))) . c ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A2, A29; ::_thesis: verum end; suppose (Initialize (IExec (I,Q,t))) . b <> (Initialize (IExec (I,Q,t))) . c ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) then H1( Initialize (IExec (I,Q,t))) = max ((abs ((Initialize (IExec (I,Q,t))) . b)),(abs ((Initialize (IExec (I,Q,t))) . c))) by A2 .= max ((abs ((IExec (I,Q,t)) . b)),(abs ((Initialize (IExec (I,Q,t))) . c))) by SCMPDS_5:15 .= max ((abs ((IExec (I,Q,t)) . b)),(abs ((IExec (I,Q,t)) . c))) by SCMPDS_5:15 .= max (((IExec (I,Q,t)) . b),(abs ((IExec (I,Q,t)) . c))) by A20, ABSVALUE:def_1 .= max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) by A20, ABSVALUE:def_1 ; hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A20, A28; ::_thesis: verum end; end; end; A30: (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) by A3, A7, A10, A16, A17, A18, A19; thus S1[ Initialize (IExec (I,Q,t))] ::_thesis: verum proof take v = Initialize (IExec (I,Q,t)); ::_thesis: ( v = Initialize (IExec (I,Q,t)) & v . b > 0 & v . c > 0 & v . (DataLoc (d,i)) = (v . b) - (v . c) ) thus v = Initialize (IExec (I,Q,t)) ; ::_thesis: ( v . b > 0 & v . c > 0 & v . (DataLoc (d,i)) = (v . b) - (v . c) ) thus ( v . b > 0 & v . c > 0 ) by A20, SCMPDS_5:15; ::_thesis: v . (DataLoc (d,i)) = (v . b) - (v . c) thus v . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) by A30, SCMPDS_5:15 .= (v . b) - ((IExec (I,Q,t)) . c) by SCMPDS_5:15 .= (v . b) - (v . c) by SCMPDS_5:15 ; ::_thesis: verum end; end; A31: for t being 0 -started State of SCMPDS st S1[t] & H1(t) = 0 holds t . (DataLoc ((s . a),i)) = 0 proof let t be 0 -started State of SCMPDS; ::_thesis: ( S1[t] & H1(t) = 0 implies t . (DataLoc ((s . a),i)) = 0 ) assume that A32: S1[t] and A33: H1(t) = 0 ; ::_thesis: t . (DataLoc ((s . a),i)) = 0 consider v being State of SCMPDS such that A34: v = t and A35: v . b > 0 and v . c > 0 and A36: v . (DataLoc (d,i)) = (v . b) - (v . c) by A32; A37: now__::_thesis:_not_t_._b_<>_t_._c assume t . b <> t . c ; ::_thesis: contradiction then t . b <> t . c ; then t . b <> t . c ; then A38: H1(t) = max ((abs (t . b)),(abs (t . c))) by A2 .= max ((abs (t . b)),(abs (t . c))) .= max ((abs (t . b)),(abs (t . c))) ; t . b > 0 by A34, A35; then abs (t . b) > 0 by COMPLEX1:47; hence contradiction by A33, A38, XXREAL_0:25; ::_thesis: verum end; thus t . (DataLoc ((s . a),i)) = (v . b) - (v . c) by A3, A34, A36 .= (t . b) - (v . c) by A34 .= (t . b) - (t . c) by A34 .= 0 by A37 ; ::_thesis: verum end; A39: S1[s] by A4, A5, A6; ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) from SCPINVAR:sch_3(A31, A39, A8); hence ( while<>0 (a,i,I) is_closed_on s,P & while<>0 (a,i,I) is_halting_on s,P ) ; ::_thesis: ( s . (DataLoc ((s . a),i)) <> 0 implies IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) assume A40: s . (DataLoc ((s . a),i)) <> 0 ; ::_thesis: IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) from SCPINVAR:sch_4(A40, A31, A39, A8); hence IExec ((while<>0 (a,i,I)),P,s) = IExec ((while<>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ; ::_thesis: verum end; begin definition func GCD-Algorithm -> Program of SCMPDS equals :: SCPINVAR:def 4 (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))); coherence (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is Program of SCMPDS ; end; :: deftheorem defines GCD-Algorithm SCPINVAR:def_4_:_ GCD-Algorithm = (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))); theorem Th21: :: SCPINVAR:17 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, b, c being Int_position for i, d being Integer st s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) 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, b, c being Int_position for i, d being Integer st s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) let s be 0 -started State of SCMPDS; ::_thesis: for I being halt-free shiftable Program of SCMPDS for a, b, c being Int_position for i, d being Integer st s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) let I be halt-free shiftable Program of SCMPDS; ::_thesis: for a, b, c being Int_position for i, d being Integer st s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) let a, b, c be Int_position; ::_thesis: for i, d being Integer st s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) holds ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) let i, d be Integer; ::_thesis: ( s . a = d & s . b > 0 & s . c > 0 & s . (DataLoc (d,i)) = (s . b) - (s . c) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) implies ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) ) set ci = DataLoc ((s . a),i); consider f being Function of (product (the_Values_of SCMPDS)),NAT such that A2: for s being State of SCMPDS holds ( ( s . b = s . c implies f . s = 0 ) & ( s . b <> s . c implies f . s = max ((abs (s . b)),(abs (s . c))) ) ) by Th6; deffunc H1( State of SCMPDS) -> Element of NAT = f . $1; set s1 = IExec ((while<>0 (a,i,I)),P,s); set ss = IExec ((while<>0 (a,i,I)),P,s); defpred S1[ set ] means ex t being 0 -started State of SCMPDS st ( t = $1 & t . b > 0 & t . c > 0 & (t . b) gcd (t . c) = (s . b) gcd (s . c) & t . (DataLoc (d,i)) = (t . b) - (t . c) ); assume that A3: s . a = d and A4: s . b > 0 and A5: s . c > 0 and A6: s . (DataLoc (d,i)) = (s . b) - (s . c) ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c & not ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ) or ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) ) assume A7: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st t . b > 0 & t . c > 0 & t . a = d & t . (DataLoc (d,i)) = (t . b) - (t . c) & t . b <> t . c holds ( (IExec (I,Q,t)) . a = d & I is_closed_on t,Q & I is_halting_on t,Q & ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) & ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) & (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) ) ; ::_thesis: ( (IExec ((while<>0 (a,i,I)),P,s)) . b = (s . b) gcd (s . c) & (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) ) A8: 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,Q,t))_._a_=_t_._a_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_H1(_Initialize_(IExec_(I,Q,t)))_<_H1(t)_&_S1[_Initialize_(IExec_(I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,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,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) ) assume that A9: S1[t] and A10: t . a = s . a and A11: t . (DataLoc ((s . a),i)) <> 0 ; ::_thesis: ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) set It = IExec (I,Q,t); set t2 = Initialize (IExec (I,Q,t)); set t1 = t; set x = (IExec (I,Q,t)) . b; set y = (IExec (I,Q,t)) . c; consider v being State of SCMPDS such that A12: v = t and A13: v . b > 0 and A14: v . c > 0 and A15: (v . b) gcd (v . c) = (s . b) gcd (s . c) and A16: v . (DataLoc (d,i)) = (v . b) - (v . c) by A9; A17: t . b > 0 by A12, A13; A18: t . c > 0 by A12, A14; A19: t . (DataLoc (d,i)) = (v . b) - (v . c) by A12, A16 .= (t . b) - (v . c) by A12 .= (t . b) - (t . c) by A12 ; then A20: ( t . b > t . c implies ( (IExec (I,Q,t)) . b = (t . b) - (t . c) & (IExec (I,Q,t)) . c = t . c ) ) by A3, A7, A10, A17, A18; A21: t . b <> t . c by A3, A11, A19; hence (IExec (I,Q,t)) . a = t . a by A3, A7, A10, A17, A18, A19; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) thus ( I is_closed_on t,Q & I is_halting_on t,Q ) by A3, A7, A10, A17, A18, A19, A21; ::_thesis: ( H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) A22: ( t . b <= t . c implies ( (IExec (I,Q,t)) . c = (t . c) - (t . b) & (IExec (I,Q,t)) . b = t . b ) ) by A3, A7, A10, A17, A18, A19, A21; A23: now__::_thesis:_(_(IExec_(I,Q,t))_._b_>_0_&_(IExec_(I,Q,t))_._c_>_0_&_((IExec_(I,Q,t))_._b)_gcd_((IExec_(I,Q,t))_._c)_=_(t_._b)_gcd_(t_._c)_&_max_(((IExec_(I,Q,t))_._b),((IExec_(I,Q,t))_._c))_<_max_((t_._b),(t_._c))_) percases ( t . b > t . c or t . b <= t . c ) ; supposeA24: t . b > t . c ; ::_thesis: ( (IExec (I,Q,t)) . b > 0 & (IExec (I,Q,t)) . c > 0 & ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) then (t . b) - (t . c) > 0 by XREAL_1:50; hence (IExec (I,Q,t)) . b > 0 by A3, A7, A10, A17, A18, A19, A24; ::_thesis: ( (IExec (I,Q,t)) . c > 0 & ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) thus (IExec (I,Q,t)) . c > 0 by A3, A7, A10, A17, A18, A19, A24; ::_thesis: ( ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) thus ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) by A17, A18, A20, A22, PREPOWER:97; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) A25: (IExec (I,Q,t)) . b = (t . b) - (t . c) by A3, A7, A10, A17, A18, A19, A24; hereby ::_thesis: verum A26: max ((t . b),(t . c)) = t . b by A24, XXREAL_0:def_10; percases ( max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b or max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c ) by XXREAL_0:16; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A18, A25, A26, XREAL_1:44; ::_thesis: verum end; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A3, A7, A10, A17, A18, A19, A24, A26; ::_thesis: verum end; end; end; end; supposeA27: t . b <= t . c ; ::_thesis: ( (IExec (I,Q,t)) . b > 0 & (IExec (I,Q,t)) . c > 0 & ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) hence (IExec (I,Q,t)) . b > 0 by A3, A7, A10, A17, A18, A19, A21; ::_thesis: ( (IExec (I,Q,t)) . c > 0 & ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) t . b < t . c by A21, A27, XXREAL_0:1; then (t . c) - (t . b) > 0 by XREAL_1:50; hence (IExec (I,Q,t)) . c > 0 by A3, A7, A10, A17, A18, A19, A21, A27; ::_thesis: ( ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) & max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) ) thus ((IExec (I,Q,t)) . b) gcd ((IExec (I,Q,t)) . c) = (t . b) gcd (t . c) by A17, A18, A20, A22, PREPOWER:97; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) A28: (IExec (I,Q,t)) . c = (t . c) - (t . b) by A3, A7, A10, A17, A18, A19, A21, A27; A29: (IExec (I,Q,t)) . b = t . b by A3, A7, A10, A17, A18, A19, A21, A27; hereby ::_thesis: verum A30: max ((t . b),(t . c)) = t . c by A27, XXREAL_0:def_10; percases ( max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c or max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b ) by XXREAL_0:16; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . c ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A17, A28, A30, XREAL_1:44; ::_thesis: verum end; suppose max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) = (IExec (I,Q,t)) . b ; ::_thesis: max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) hence max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) < max ((t . b),(t . c)) by A21, A27, A29, A30, XXREAL_0:1; ::_thesis: verum end; end; end; end; end; end; thus H1( Initialize (IExec (I,Q,t))) < H1(t) ::_thesis: S1[ Initialize (IExec (I,Q,t))] proof t . b <> t . c by A21; then t . b <> t . c ; then A31: H1(t) = max ((abs (t . b)),(abs (t . c))) by A2 .= max ((abs (t . b)),(abs (t . c))) .= max ((abs (t . b)),(abs (t . c))) .= max ((t . b),(abs (t . c))) by A17, ABSVALUE:def_1 .= max ((t . b),(t . c)) by A18, ABSVALUE:def_1 ; then H1(t) >= t . b by XXREAL_0:25; then A32: H1(t) > 0 by A17, XXREAL_0:2; percases ( (Initialize (IExec (I,Q,t))) . b = (Initialize (IExec (I,Q,t))) . c or (Initialize (IExec (I,Q,t))) . b <> (Initialize (IExec (I,Q,t))) . c ) ; suppose (Initialize (IExec (I,Q,t))) . b = (Initialize (IExec (I,Q,t))) . c ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A2, A32; ::_thesis: verum end; suppose (Initialize (IExec (I,Q,t))) . b <> (Initialize (IExec (I,Q,t))) . c ; ::_thesis: H1( Initialize (IExec (I,Q,t))) < H1(t) then H1( Initialize (IExec (I,Q,t))) = max ((abs ((Initialize (IExec (I,Q,t))) . b)),(abs ((Initialize (IExec (I,Q,t))) . c))) by A2 .= max ((abs ((IExec (I,Q,t)) . b)),(abs ((Initialize (IExec (I,Q,t))) . c))) by SCMPDS_5:15 .= max ((abs ((IExec (I,Q,t)) . b)),(abs ((IExec (I,Q,t)) . c))) by SCMPDS_5:15 .= max (((IExec (I,Q,t)) . b),(abs ((IExec (I,Q,t)) . c))) by A23, ABSVALUE:def_1 .= max (((IExec (I,Q,t)) . b),((IExec (I,Q,t)) . c)) by A23, ABSVALUE:def_1 ; hence H1( Initialize (IExec (I,Q,t))) < H1(t) by A23, A31; ::_thesis: verum end; end; end; A33: (IExec (I,Q,t)) . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) by A3, A7, A10, A17, A18, A19, A21; thus S1[ Initialize (IExec (I,Q,t))] ::_thesis: verum proof take u = Initialize (IExec (I,Q,t)); ::_thesis: ( u = Initialize (IExec (I,Q,t)) & u . b > 0 & u . c > 0 & (u . b) gcd (u . c) = (s . b) gcd (s . c) & u . (DataLoc (d,i)) = (u . b) - (u . c) ) thus u = Initialize (IExec (I,Q,t)) ; ::_thesis: ( u . b > 0 & u . c > 0 & (u . b) gcd (u . c) = (s . b) gcd (s . c) & u . (DataLoc (d,i)) = (u . b) - (u . c) ) thus ( u . b > 0 & u . c > 0 ) by A23, SCMPDS_5:15; ::_thesis: ( (u . b) gcd (u . c) = (s . b) gcd (s . c) & u . (DataLoc (d,i)) = (u . b) - (u . c) ) thus (u . b) gcd (u . c) = ((IExec (I,Q,t)) . b) gcd (u . c) by SCMPDS_5:15 .= (t . b) gcd (t . c) by A23, SCMPDS_5:15 .= (v . b) gcd (t . c) by A12 .= (s . b) gcd (s . c) by A12, A15 ; ::_thesis: u . (DataLoc (d,i)) = (u . b) - (u . c) thus u . (DataLoc (d,i)) = ((IExec (I,Q,t)) . b) - ((IExec (I,Q,t)) . c) by A33, SCMPDS_5:15 .= (u . b) - ((IExec (I,Q,t)) . c) by SCMPDS_5:15 .= (u . b) - (u . c) by SCMPDS_5:15 ; ::_thesis: verum end; end; A34: for t being 0 -started State of SCMPDS st S1[t] holds ( H1(t) = 0 iff t . (DataLoc ((s . a),i)) = 0 ) proof let t be 0 -started State of SCMPDS; ::_thesis: ( S1[t] implies ( H1(t) = 0 iff t . (DataLoc ((s . a),i)) = 0 ) ) assume S1[t] ; ::_thesis: ( H1(t) = 0 iff t . (DataLoc ((s . a),i)) = 0 ) then consider v being State of SCMPDS such that A35: v = t and A36: v . b > 0 and v . c > 0 and (v . b) gcd (v . c) = (s . b) gcd (s . c) and A37: v . (DataLoc (d,i)) = (v . b) - (v . c) ; A38: t . (DataLoc ((s . a),i)) = (v . b) - (v . c) by A3, A35, A37 .= (t . b) - (v . c) by A35 .= (t . b) - (t . c) by A35 ; hereby ::_thesis: ( t . (DataLoc ((s . a),i)) = 0 implies H1(t) = 0 ) assume A39: H1(t) = 0 ; ::_thesis: t . (DataLoc ((s . a),i)) = 0 now__::_thesis:_not_t_._b_<>_t_._c assume t . b <> t . c ; ::_thesis: contradiction then t . b <> t . c ; then t . b <> t . c ; then A40: H1(t) = max ((abs (t . b)),(abs (t . c))) by A2 .= max ((abs (t . b)),(abs (t . c))) .= max ((abs (t . b)),(abs (t . c))) ; t . b > 0 by A35, A36; then abs (t . b) > 0 by COMPLEX1:47; hence contradiction by A39, A40, XXREAL_0:25; ::_thesis: verum end; hence t . (DataLoc ((s . a),i)) = 0 by A38; ::_thesis: verum end; thus ( t . (DataLoc ((s . a),i)) = 0 implies H1(t) = 0 ) by A2, A38; ::_thesis: verum end; A41: S1[s] by A4, A5, A6; A42: ( H1( Initialize (IExec ((while<>0 (a,i,I)),P,s))) = 0 & S1[ Initialize (IExec ((while<>0 (a,i,I)),P,s))] ) from SCPINVAR:sch_5(A34, A41, A8); then consider w being 0 -started State of SCMPDS such that A43: w = Initialize (IExec ((while<>0 (a,i,I)),P,s)) and A44: w . b > 0 and w . c > 0 and A45: (w . b) gcd (w . c) = (s . b) gcd (s . c) and A46: w . (DataLoc (d,i)) = (w . b) - (w . c) ; CI: (Initialize (IExec ((while<>0 (a,i,I)),P,s))) . (DataLoc ((s . a),i)) = (IExec ((while<>0 (a,i,I)),P,s)) . (DataLoc ((s . a),i)) by SCMPDS_5:15; B: (Initialize (IExec ((while<>0 (a,i,I)),P,s))) . b = (IExec ((while<>0 (a,i,I)),P,s)) . b by SCMPDS_5:15; C: (Initialize (IExec ((while<>0 (a,i,I)),P,s))) . c = (IExec ((while<>0 (a,i,I)),P,s)) . c by SCMPDS_5:15; A47: (w . b) - (w . c) = (IExec ((while<>0 (a,i,I)),P,s)) . (DataLoc ((s . a),i)) by A3, A43, A46, SCMPDS_5:15 .= 0 by A34, A42, CI ; then A48: abs (w . b) = (abs (w . b)) gcd (abs (w . c)) by NAT_D:32 .= (s . b) gcd (s . c) by A45, INT_2:34 ; thus (IExec ((while<>0 (a,i,I)),P,s)) . b = (IExec ((while<>0 (a,i,I)),P,s)) . b .= (s . b) gcd (s . c) by A43, A44, A48, B, ABSVALUE:def_1 ; ::_thesis: (IExec ((while<>0 (a,i,I)),P,s)) . c = (s . b) gcd (s . c) thus (IExec ((while<>0 (a,i,I)),P,s)) . c = (IExec ((while<>0 (a,i,I)),P,s)) . c .= (s . b) gcd (s . c) by A43, A44, A47, A48, C, ABSVALUE:def_1 ; ::_thesis: verum end; set i1 = GBP := 0; set i2 = (GBP,3) := (GBP,1); set i3 = SubFrom (GBP,3,GBP,2); set j1 = Load (SubFrom (GBP,1,GBP,2)); set j2 = Load (SubFrom (GBP,2,GBP,1)); set IF = if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1)))); set k1 = (GBP,3) := (GBP,1); set k2 = SubFrom (GBP,3,GBP,2); set WB = ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)); set WH = while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))); Lm12: card (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) = 6 proof thus card (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) = (card ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1)))) + 1 by SCMP_GCD:4 .= ((card (if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1)))))) + 1) + 1 by SCMP_GCD:4 .= ((((card (Load (SubFrom (GBP,1,GBP,2)))) + (card (Load (SubFrom (GBP,2,GBP,1))))) + 2) + 1) + 1 by SCMPDS_6:65 .= (((1 + (card (Load (SubFrom (GBP,2,GBP,1))))) + 2) + 1) + 1 by COMPOS_1:54 .= (((1 + 1) + 2) + 1) + 1 by COMPOS_1:54 .= 6 ; ::_thesis: verum end; Lm13: card (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) = 9 proof thus card (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) = 6 + 3 by Lm12, Th12 .= 9 ; ::_thesis: verum end; theorem :: SCPINVAR:18 card GCD-Algorithm = 12 proof thus card GCD-Algorithm = (card (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) + (card (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))) by AFINSQ_1:17 .= ((card ((GBP := 0) ';' ((GBP,3) := (GBP,1)))) + 1) + (card (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))) by SCMP_GCD:4 .= (2 + 1) + 9 by Lm13, SCMP_GCD:5 .= 12 ; ::_thesis: verum end; Lm14: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . GBP = 0 holds ( ( s . (intpos 3) > 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = s . (intpos 2) ) ) & ( s . (intpos 3) <= 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = s . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . GBP = 0 holds ( ( s . (intpos 3) > 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = s . (intpos 2) ) ) & ( s . (intpos 3) <= 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = s . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . GBP = 0 implies ( ( s . (intpos 3) > 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = s . (intpos 2) ) ) & ( s . (intpos 3) <= 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = s . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) ) set s1 = IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s); set s2 = IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s); set a = GBP ; set t0 = s; set Q0 = P; A4: now__::_thesis:_(_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._GBP_=_0_implies_(_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),P,s))_._GBP_=_0_&_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),P,s))_._(intpos_1)_=_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._(intpos_1)_&_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),P,s))_._(intpos_2)_=_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._(intpos_2)_&_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),P,s))_._(intpos_3)_=_((IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),P,s))_._(intpos_1))_-_((IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),P,s))_._(intpos_2))_)_) assume A5: (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 ; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) then A6: DataLoc (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; A7: (IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 3) = (Exec (((GBP,3) := (GBP,1)),(IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (DataLoc (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP),1)) by A6, SCMPDS_2:47 .= (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos (0 + 1)) by A5, SCMP_GCD:1 ; 1 <> abs (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP) + 3) by A5, ABSVALUE:def_1; then A8: intpos 1 <> DataLoc (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP),3) by XTUPLE_0:1; A9: (IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 1) = (Exec (((GBP,3) := (GBP,1)),(IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) by A8, SCMPDS_2:47 ; 2 <> abs (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP) + 3) by A5, ABSVALUE:def_1; then A10: intpos 2 <> DataLoc (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP),3) by XTUPLE_0:1; A11: (IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 2) = (Exec (((GBP,3) := (GBP,1)),(IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) by A10, SCMPDS_2:47 ; 0 <> abs (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP) + 3) by A5, ABSVALUE:def_1; then A12: GBP <> DataLoc (((IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP),3) by XTUPLE_0:1; A13: (IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP = (Exec (((GBP,3) := (GBP,1)),(IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A5, A12, SCMPDS_2:47 ; then A14: DataLoc (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; 0 <> abs (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP) + 3) by A13, ABSVALUE:def_1; then A15: GBP <> DataLoc (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) by XTUPLE_0:1; thus (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A13, A15, SCMPDS_2:50 ; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) 1 <> abs (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP) + 3) by A13, ABSVALUE:def_1; then A16: intpos 1 <> DataLoc (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) by XTUPLE_0:1; thus A17: (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) by A9, A16, SCMPDS_2:50 ; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) 2 <> abs (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP) + 3) by A13, ABSVALUE:def_1; then A18: intpos 2 <> DataLoc (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) by XTUPLE_0:1; thus A19: (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) by A11, A18, SCMPDS_2:50 ; ::_thesis: (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) thus (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= ((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 3)) - ((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . (DataLoc (((IExec (((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),2))) by A14, SCMPDS_2:50 .= ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) by A13, A11, A7, A17, A19, SCMP_GCD:1 ; ::_thesis: verum end; set s0 = s; set m1 = SubFrom (GBP,1,GBP,2); set m2 = SubFrom (GBP,2,GBP,1); assume A20: s . GBP = 0 ; ::_thesis: ( ( s . (intpos 3) > 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = s . (intpos 2) ) ) & ( s . (intpos 3) <= 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = s . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) then A21: s . GBP = 0 ; A22: DataLoc ((s . GBP),3) = intpos (0 + 3) by A20, SCMP_GCD:1; A24: now__::_thesis:_(_s_._(intpos_3)_>_0_implies_(_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._GBP_=_0_&_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._(intpos_1)_=_(s_._(intpos_1))_-_(s_._(intpos_2))_&_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._(intpos_2)_=_s_._(intpos_2)_)_) 2 <> abs ((s . GBP) + 1) by A21, ABSVALUE:def_1; then A25: intpos 2 <> DataLoc ((s . GBP),1) by XTUPLE_0:1; 0 <> abs ((s . GBP) + 1) by A21, ABSVALUE:def_1; then A26: GBP <> DataLoc ((s . GBP),1) by XTUPLE_0:1; assume A27: s . (intpos 3) > 0 ; ::_thesis: ( (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 & (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = s . (intpos 2) ) hence (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = (IExec ((Load (SubFrom (GBP,1,GBP,2))),P,s)) . GBP by A22, SCMPDS_6:73 .= (Exec ((SubFrom (GBP,1,GBP,2)),s)) . GBP by SCMPDS_5:40 .= 0 by A21, A26, SCMPDS_2:50 ; ::_thesis: ( (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = s . (intpos 2) ) A28: DataLoc ((s . GBP),1) = intpos (0 + 1) by A21, SCMP_GCD:1; thus (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = (IExec ((Load (SubFrom (GBP,1,GBP,2))),P,s)) . (intpos 1) by A22, A27, SCMPDS_6:73 .= (Exec ((SubFrom (GBP,1,GBP,2)),s)) . (intpos 1) by SCMPDS_5:40 .= (s . (intpos 1)) - (s . (DataLoc ((s . GBP),2))) by A28, SCMPDS_2:50 .= (s . (intpos 1)) - (s . (intpos (0 + 2))) by A21, SCMP_GCD:1 .= (s . (intpos 1)) - (s . (intpos 2)) ; ::_thesis: (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = s . (intpos 2) thus (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = (IExec ((Load (SubFrom (GBP,1,GBP,2))),P,s)) . (intpos 2) by A22, A27, SCMPDS_6:73 .= (Exec ((SubFrom (GBP,1,GBP,2)),s)) . (intpos 2) by SCMPDS_5:40 .= s . (intpos 2) by A25, SCMPDS_2:50 ; ::_thesis: verum end; hence ( s . (intpos 3) > 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (s . (intpos 1)) - (s . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = s . (intpos 2) ) ) by A4; ::_thesis: ( ( s . (intpos 3) <= 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = s . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) A30: now__::_thesis:_(_s_._(intpos_3)_<=_0_implies_(_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._GBP_=_0_&_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._(intpos_2)_=_(s_._(intpos_2))_-_(s_._(intpos_1))_&_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._(intpos_1)_=_s_._(intpos_1)_)_) 1 <> abs ((s . GBP) + 2) by A21, ABSVALUE:def_1; then A31: intpos 1 <> DataLoc ((s . GBP),2) by XTUPLE_0:1; 0 <> abs ((s . GBP) + 2) by A21, ABSVALUE:def_1; then A32: GBP <> DataLoc ((s . GBP),2) by XTUPLE_0:1; assume A33: s . (intpos 3) <= 0 ; ::_thesis: ( (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 & (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = s . (intpos 1) ) hence (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = (IExec ((Load (SubFrom (GBP,2,GBP,1))),P,s)) . GBP by A22, SCMPDS_6:74 .= (Exec ((SubFrom (GBP,2,GBP,1)),s)) . GBP by SCMPDS_5:40 .= 0 by A21, A32, SCMPDS_2:50 ; ::_thesis: ( (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = s . (intpos 1) ) A34: DataLoc ((s . GBP),2) = intpos (0 + 2) by A21, SCMP_GCD:1; thus (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 2) = (IExec ((Load (SubFrom (GBP,2,GBP,1))),P,s)) . (intpos 2) by A22, A33, SCMPDS_6:74 .= (Exec ((SubFrom (GBP,2,GBP,1)),s)) . (intpos 2) by SCMPDS_5:40 .= (s . (intpos 2)) - (s . (DataLoc ((s . GBP),1))) by A34, SCMPDS_2:50 .= (s . (intpos 2)) - (s . (intpos (0 + 1))) by A21, SCMP_GCD:1 .= (s . (intpos 2)) - (s . (intpos 1)) ; ::_thesis: (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = s . (intpos 1) thus (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . (intpos 1) = (IExec ((Load (SubFrom (GBP,2,GBP,1))),P,s)) . (intpos 1) by A22, A33, SCMPDS_6:74 .= (Exec ((SubFrom (GBP,2,GBP,1)),s)) . (intpos 1) by SCMPDS_5:40 .= s . (intpos 1) by A31, SCMPDS_2:50 ; ::_thesis: verum end; hence ( s . (intpos 3) <= 0 implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (s . (intpos 2)) - (s . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = s . (intpos 1) ) ) by A4; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) now__::_thesis:_(IExec_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1))))),P,s))_._GBP_=_0 percases ( s . (intpos 3) > 0 or s . (intpos 3) <= 0 ) ; suppose s . (intpos 3) > 0 ; ::_thesis: (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 hence (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 by A24; ::_thesis: verum end; suppose s . (intpos 3) <= 0 ; ::_thesis: (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 hence (IExec ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))),P,s)) . GBP = 0 by A30; ::_thesis: verum end; end; end; hence ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = 0 & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) ) by A4; ::_thesis: verum end; Lm15: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . GBP = 0 & s . (intpos 1) > 0 & s . (intpos 2) > 0 & s . (intpos 3) = (s . (intpos 1)) - (s . (intpos 2)) holds ( (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on s,P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on s,P ) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started State of SCMPDS st s . GBP = 0 & s . (intpos 1) > 0 & s . (intpos 2) > 0 & s . (intpos 3) = (s . (intpos 1)) - (s . (intpos 2)) holds ( (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on s,P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on s,P ) set a = GBP ; let s be 0 -started State of SCMPDS; ::_thesis: ( s . GBP = 0 & s . (intpos 1) > 0 & s . (intpos 2) > 0 & s . (intpos 3) = (s . (intpos 1)) - (s . (intpos 2)) implies ( (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on s,P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on s,P ) ) A1: DataLoc (0,3) = intpos (0 + 3) by SCMP_GCD:1; A2: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_t_._(intpos_1)_>_0_&_t_._(intpos_2)_>_0_&_t_._GBP_=_0_&_t_._(DataLoc_(0,3))_=_(t_._(intpos_1))_-_(t_._(intpos_2))_&_t_._(intpos_1)_<>_t_._(intpos_2)_holds_ (_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._GBP_=_0_&_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))_is_closed_on_t,Q_&_((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))_is_halting_on_t,Q_&_(_t_._(intpos_1)_>_t_._(intpos_2)_implies_(_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(intpos_1)_=_(t_._(intpos_1))_-_(t_._(intpos_2))_&_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(intpos_2)_=_t_._(intpos_2)_)_)_&_(_t_._(intpos_1)_<=_t_._(intpos_2)_implies_(_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(intpos_2)_=_(t_._(intpos_2))_-_(t_._(intpos_1))_&_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(intpos_1)_=_t_._(intpos_1)_)_)_&_(IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(DataLoc_(0,3))_=_((IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(intpos_1))_-_((IExec_((((if>0_(GBP,3,(Load_(SubFrom_(GBP,1,GBP,2))),(Load_(SubFrom_(GBP,2,GBP,1)))))_';'_((GBP,3)_:=_(GBP,1)))_';'_(SubFrom_(GBP,3,GBP,2))),Q,t))_._(intpos_2))_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st t . (intpos 1) > 0 & t . (intpos 2) > 0 & t . GBP = 0 & t . (DataLoc (0,3)) = (t . (intpos 1)) - (t . (intpos 2)) & t . (intpos 1) <> t . (intpos 2) holds ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . GBP = 0 & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_closed_on t,Q & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_halting_on t,Q & ( t . (intpos 1) > t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) ) & ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( t . (intpos 1) > 0 & t . (intpos 2) > 0 & t . GBP = 0 & t . (DataLoc (0,3)) = (t . (intpos 1)) - (t . (intpos 2)) & t . (intpos 1) <> t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . GBP = 0 & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_closed_on t,Q & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_halting_on t,Q & ( t . (intpos 1) > t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) ) & ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) ) assume that t . (intpos 1) > 0 and t . (intpos 2) > 0 and A3: t . GBP = 0 and A4: t . (DataLoc (0,3)) = (t . (intpos 1)) - (t . (intpos 2)) and t . (intpos 1) <> t . (intpos 2) ; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . GBP = 0 & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_closed_on t,Q & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_halting_on t,Q & ( t . (intpos 1) > t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) ) & ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) thus (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . GBP = 0 by A3, Lm14; ::_thesis: ( ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_closed_on t,Q & ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_halting_on t,Q & ( t . (intpos 1) > t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) ) & ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) thus ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_closed_on t,Q by SCMPDS_6:20; ::_thesis: ( ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_halting_on t,Q & ( t . (intpos 1) > t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) ) & ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) thus ((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)) is_halting_on t,Q by SCMPDS_6:21; ::_thesis: ( ( t . (intpos 1) > t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) ) & ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) hereby ::_thesis: ( ( t . (intpos 1) <= t . (intpos 2) implies ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) ) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) ) assume t . (intpos 1) > t . (intpos 2) ; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) then t . (intpos 3) > 0 by A1, A4, XREAL_1:50; hence ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = (t . (intpos 1)) - (t . (intpos 2)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = t . (intpos 2) ) by A3, Lm14; ::_thesis: verum end; hereby ::_thesis: (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) assume t . (intpos 1) <= t . (intpos 2) ; ::_thesis: ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) then t . (intpos 3) <= 0 by A1, A4, XREAL_1:47; hence ( (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2) = (t . (intpos 2)) - (t . (intpos 1)) & (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1) = t . (intpos 1) ) by A3, Lm14; ::_thesis: verum end; thus (IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (DataLoc (0,3)) = ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 1)) - ((IExec ((((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),Q,t)) . (intpos 2)) by A1, A3, Lm14; ::_thesis: verum end; assume A5: ( s . GBP = 0 & s . (intpos 1) > 0 & s . (intpos 2) > 0 & s . (intpos 3) = (s . (intpos 1)) - (s . (intpos 2)) ) ; ::_thesis: ( (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on s,P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on s,P ) hence ( (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) ) by A1, A2, Th21; ::_thesis: ( while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on s,P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on s,P ) thus ( while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on s,P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on s,P ) by A5, A1, A2, Lm12, Th20; ::_thesis: verum end; set GA = (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))); Lm16: for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS st s . (intpos 1) > 0 & s . (intpos 2) > 0 holds ( (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) 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 1) > 0 & s . (intpos 2) > 0 holds ( (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_halting_on s,P ) let s be 0 -started State of SCMPDS; ::_thesis: ( s . (intpos 1) > 0 & s . (intpos 2) > 0 implies ( (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_halting_on s,P ) ) I: Initialize s = s by MEMSTR_0:44; assume A1: ( s . (intpos 1) > 0 & s . (intpos 2) > 0 ) ; ::_thesis: ( (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) & (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_halting_on s,P ) set t0 = s; set Q0 = P; set t1 = IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s); set t2 = IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s); set Q1 = P; set t3 = Exec ((GBP := 0),s); set a = GBP ; A5: (Exec ((GBP := 0),s)) . (intpos 1) = s . (intpos 1) by AMI_3:10, SCMPDS_2:45 .= s . (intpos 1) ; A6: (Exec ((GBP := 0),s)) . GBP = 0 by SCMPDS_2:45; then A7: DataLoc (((Exec ((GBP := 0),s)) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; 1 <> abs (((Exec ((GBP := 0),s)) . GBP) + 3) by A6, ABSVALUE:def_1; then A8: intpos 1 <> DataLoc (((Exec ((GBP := 0),s)) . GBP),3) by XTUPLE_0:1; A9: (IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 1) = (Exec (((GBP,3) := (GBP,1)),(Exec ((GBP := 0),s)))) . (intpos 1) by SCMPDS_5:42 .= s . (intpos 1) by A5, A8, SCMPDS_2:47 ; A10: (Exec ((GBP := 0),s)) . (intpos 2) = s . (intpos 2) by AMI_3:10, SCMPDS_2:45 .= s . (intpos 2) ; A11: (IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 3) = (Exec (((GBP,3) := (GBP,1)),(Exec ((GBP := 0),s)))) . (intpos 3) by SCMPDS_5:42 .= (Exec ((GBP := 0),s)) . (DataLoc (((Exec ((GBP := 0),s)) . GBP),1)) by A7, SCMPDS_2:47 .= s . (intpos 1) by A6, A5, SCMP_GCD:1 ; 2 <> abs (((Exec ((GBP := 0),s)) . GBP) + 3) by A6, ABSVALUE:def_1; then A12: intpos 2 <> DataLoc (((Exec ((GBP := 0),s)) . GBP),3) by XTUPLE_0:1; A13: (IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 2) = (Exec (((GBP,3) := (GBP,1)),(Exec ((GBP := 0),s)))) . (intpos 2) by SCMPDS_5:42 .= s . (intpos 2) by A10, A12, SCMPDS_2:47 ; 0 <> abs (((Exec ((GBP := 0),s)) . GBP) + 3) by A6, ABSVALUE:def_1; then A14: GBP <> DataLoc (((Exec ((GBP := 0),s)) . GBP),3) by XTUPLE_0:1; A15: (IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP = (Exec (((GBP,3) := (GBP,1)),(Exec ((GBP := 0),s)))) . GBP by SCMPDS_5:42 .= 0 by A6, A14, SCMPDS_2:47 ; then A16: DataLoc (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) = intpos (0 + 3) by SCMP_GCD:1; 0 <> abs (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP) + 3) by A15, ABSVALUE:def_1; then A17: GBP <> DataLoc (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) by XTUPLE_0:1; A18: (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)))) . GBP by SCMPDS_5:41 .= 0 by A15, A17, SCMPDS_2:50 ; 1 <> abs (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP) + 3) by A15, ABSVALUE:def_1; then A19: intpos 1 <> DataLoc (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) by XTUPLE_0:1; A20: (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)))) . (intpos 1) by SCMPDS_5:41 .= s . (intpos 1) by A9, A19, SCMPDS_2:50 ; 2 <> abs (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP) + 3) by A15, ABSVALUE:def_1; then A21: intpos 2 <> DataLoc (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),3) by XTUPLE_0:1; A22: (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)))) . (intpos 2) by SCMPDS_5:41 .= s . (intpos 2) by A13, A21, SCMPDS_2:50 ; X1: (Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))) . (intpos 1) = (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1) by SCMPDS_5:15; X2: (Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))) . (intpos 2) = (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2) by SCMPDS_5:15; A23: (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) = (Exec ((SubFrom (GBP,3,GBP,2)),(IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)))) . (intpos 3) by SCMPDS_5:41 .= ((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . (intpos 3)) - ((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . (DataLoc (((IExec (((GBP := 0) ';' ((GBP,3) := (GBP,1))),P,s)) . GBP),2))) by A16, SCMPDS_2:50 .= ((IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) - ((IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) by A15, A13, A11, A20, A22, SCMP_GCD:1 ; X0: (Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))) . GBP = (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . GBP by SCMPDS_5:15; X3: (Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))) . (intpos 3) = (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 3) by SCMPDS_5:15; B25: ( while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)),P & while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)),P ) by A1, A18, A20, A22, Lm15, X1, X2, A23, X0, X3; A24: while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s),P proof for k being Element of NAT holds IC (Comput ((P +* (stop (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))))),(Initialize (Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)))),k)) in dom (stop (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))) by B25, SCMPDS_6:def_2; hence while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_closed_on IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s),P by SCMPDS_6:def_2; ::_thesis: verum end; A25: while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s),P proof P +* (stop (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))) halts_on Initialize (Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))) by B25, SCMPDS_6:def_3; hence while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))) is_halting_on IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s),P by SCMPDS_6:def_3; ::_thesis: verum end; (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,(Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))))) . (intpos 1) = ((IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) gcd ((IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) by A1, A18, A20, A22, A23, Lm15, X1, X2, X3, X0; hence (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 1) = (s . (intpos 1)) gcd (s . (intpos 2)) by A20, A22, A24, A25, SCPISORT:7; ::_thesis: ( (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_halting_on s,P ) (IExec ((while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))),P,(Initialize (IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s))))) . (intpos 2) = ((IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 1)) gcd ((IExec ((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))),P,s)) . (intpos 2)) by A1, A18, A20, A22, A23, Lm15, X1, X2, X3, X0; hence (IExec (((((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2)))))),P,s)) . (intpos 2) = (s . (intpos 1)) gcd (s . (intpos 2)) by A20, A22, A24, A25, SCPISORT:7; ::_thesis: ( (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_halting_on s,P ) thus ( (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_closed_on s,P & (((GBP := 0) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))) ';' (while<>0 (GBP,3,(((if>0 (GBP,3,(Load (SubFrom (GBP,1,GBP,2))),(Load (SubFrom (GBP,2,GBP,1))))) ';' ((GBP,3) := (GBP,1))) ';' (SubFrom (GBP,3,GBP,2))))) is_halting_on s,P ) by A24, I, A25, SCPISORT:9; ::_thesis: verum end; theorem :: SCPINVAR:19 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for x, y being Integer st s . (intpos 1) = x & s . (intpos 2) = y & x > 0 & y > 0 holds ( (IExec (GCD-Algorithm,P,s)) . (intpos 1) = x gcd y & (IExec (GCD-Algorithm,P,s)) . (intpos 2) = x gcd y & GCD-Algorithm is_closed_on s,P & GCD-Algorithm is_halting_on s,P ) by Lm16;