:: SCMPDS_8 semantic presentation begin set A = NAT ; set D = SCM-Data-Loc ; theorem :: SCMPDS_8:1 for a being Int_position ex i being Element of NAT st a = intpos i proof let a be Int_position; ::_thesis: ex i being Element of NAT st a = intpos i a in SCM-Data-Loc by AMI_2:def_16; then consider x, y being set such that A1: x in {1} and A2: y in NAT and A3: a = [x,y] by ZFMISC_1:84; reconsider k = y as Element of NAT by A2; take k ; ::_thesis: a = intpos k thus intpos k = dl. k by SCMP_GCD:def_1 .= a by A1, A3, TARSKI:def_1 ; ::_thesis: verum end; theorem :: SCMPDS_8:2 canceled; theorem Th3: :: SCMPDS_8:3 for t being State of SCMPDS for i being Instruction of SCMPDS st InsCode i in {0,4,5,6,14} holds Initialize t = Initialize (Exec (i,t)) proof let t be State of SCMPDS; ::_thesis: for i being Instruction of SCMPDS st InsCode i in {0,4,5,6,14} holds Initialize t = Initialize (Exec (i,t)) let i be Instruction of SCMPDS; ::_thesis: ( InsCode i in {0,4,5,6,14} implies Initialize t = Initialize (Exec (i,t)) ) assume InsCode i in {0,4,5,6,14} ; ::_thesis: Initialize t = Initialize (Exec (i,t)) then DataPart (Exec (i,t)) = DataPart t by SCMPDS_7:7; hence Initialize t = Initialize (Exec (i,t)) by MEMSTR_0:80; ::_thesis: verum end; theorem :: SCMPDS_8:4 canceled; theorem Th5: :: SCMPDS_8:5 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 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; begin definition canceled; 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 :: SCMPDS_8:def 2 (((a,i) >=0_goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 1))); coherence (((a,i) >=0_goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 1))) is Program of SCMPDS ; end; :: deftheorem SCMPDS_8:def_1_:_ canceled; :: deftheorem defines while<0 SCMPDS_8:def_2_:_ 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 ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 1))); 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 ((card I) + 2); set PF = (Load ((a,i) >=0_goto ((card I) + 2))) ';' I; A1: (Load ((a,i) >=0_goto ((card I) + 2))) ';' I = ((a,i) >=0_goto ((card I) + 2)) ';' I by SCMPDS_4:def_2; then card ((Load ((a,i) >=0_goto ((card I) + 2))) ';' I) = (card I) + 1 by SCMPDS_6:6; then (card ((Load ((a,i) >=0_goto ((card I) + 2))) ';' I)) + (- ((card I) + 1)) = 0 ; hence while<0 (a,i,I) is shiftable by A1, 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 ; ; end; theorem Th6: :: SCMPDS_8:6 for a being Int_position for i being Integer for I being Program of SCMPDS holds card (while<0 (a,i,I)) = (card I) + 2 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) + 2 let i be Integer; ::_thesis: for I being Program of SCMPDS holds card (while<0 (a,i,I)) = (card I) + 2 let I be Program of SCMPDS; ::_thesis: card (while<0 (a,i,I)) = (card I) + 2 set i1 = (a,i) >=0_goto ((card I) + 2); set I4 = ((a,i) >=0_goto ((card I) + 2)) ';' I; thus card (while<0 (a,i,I)) = (card (((a,i) >=0_goto ((card I) + 2)) ';' I)) + 1 by SCMP_GCD:4 .= ((card I) + 1) + 1 by SCMPDS_6:6 .= (card I) + 2 ; ::_thesis: verum end; Lm1: 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) + 3 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) + 3 let i be Integer; ::_thesis: for I being Program of SCMPDS holds card (stop (while<0 (a,i,I))) = (card I) + 3 let I be Program of SCMPDS; ::_thesis: card (stop (while<0 (a,i,I))) = (card I) + 3 thus card (stop (while<0 (a,i,I))) = (card (while<0 (a,i,I))) + 1 by COMPOS_1:55 .= ((card I) + 2) + 1 by Th6 .= (card I) + 3 ; ::_thesis: verum end; theorem Th7: :: SCMPDS_8:7 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) + 2 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) + 2 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) + 2 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) + 2 iff m in dom (while<0 (a,i,I)) ) let I be Program of SCMPDS; ::_thesis: ( m < (card I) + 2 iff m in dom (while<0 (a,i,I)) ) card (while<0 (a,i,I)) = (card I) + 2 by Th6; hence ( m < (card I) + 2 iff m in dom (while<0 (a,i,I)) ) by AFINSQ_1:66; ::_thesis: verum end; theorem Th8: :: SCMPDS_8:8 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 ((card I) + 2) & (while<0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) 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 ((card I) + 2) & (while<0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) let i be Integer; ::_thesis: for I being Program of SCMPDS holds ( (while<0 (a,i,I)) . 0 = (a,i) >=0_goto ((card I) + 2) & (while<0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) let I be Program of SCMPDS; ::_thesis: ( (while<0 (a,i,I)) . 0 = (a,i) >=0_goto ((card I) + 2) & (while<0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) set i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); set I4 = ((a,i) >=0_goto ((card I) + 2)) ';' I; set J5 = I ';' (goto (- ((card I) + 1))); set FLOOP = while<0 (a,i,I); while<0 (a,i,I) = ((a,i) >=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; hence (while<0 (a,i,I)) . 0 = (a,i) >=0_goto ((card I) + 2) by SCMPDS_6:7; ::_thesis: (while<0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) card (((a,i) >=0_goto ((card I) + 2)) ';' I) = (card I) + 1 by SCMPDS_6:6; hence (while<0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) by SCMP_GCD:6; ::_thesis: verum end; theorem Th9: :: SCMPDS_8:9 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 d1 = 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 i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); set WHL = while<0 (a,i,I); set pWHL = 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 P4 = P +* (stop (while<0 (a,i,I))); A3: IC (Initialize s) = 0 by MEMSTR_0:def_11; A4: 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 A5: (Initialize s) . (DataLoc (((Initialize s) . a),i)) = (Initialize s) . (DataLoc ((s . a),i)) by FUNCT_4:11 .= s . (DataLoc ((s . a),i)) by A4, FUNCT_4:11 ; A6: while<0 (a,i,I) = ((a,i) >=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; 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 ((card I) + 2)),(Initialize s)) by A6, SCMPDS_6:11 ; then A7: IC (Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize s),1)) = ICplusConst ((Initialize s),((card I) + 2)) by A1, A5, SCMPDS_2:57 .= 0 + ((card I) + 2) by A3, SCMPDS_6:12 ; A8: card (while<0 (a,i,I)) = (card I) + 2 by Th6; then A9: (card I) + 2 in dom (stop (while<0 (a,i,I))) by COMPOS_1:64; stop (while<0 (a,i,I)) c= P +* (stop (while<0 (a,i,I))) by FUNCT_4:25; then (P +* (stop (while<0 (a,i,I)))) . ((card I) + 2) = (stop (while<0 (a,i,I))) . ((card I) + 2) by A9, GRFUNC_1:2 .= halt SCMPDS by A8, COMPOS_1:64 ; then A11: CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize s),1))) = halt SCMPDS by A7, PBOOLE:143; 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))) percases ( 0 < k or k = 0 ) ; suppose 0 < k ; ::_thesis: IC (Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize s),b1)) in dom (stop (while<0 (a,i,I))) then 1 + 0 <= 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 A9, A7, A11, EXTPRO_1:5; ::_thesis: verum end; 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 A3, COMPOS_1:36; ::_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 A11, EXTPRO_1:29; hence while<0 (a,i,I) is_halting_on s,P by SCMPDS_6:def_3; ::_thesis: verum end; theorem Th10: :: SCMPDS_8:10 for P being Instruction-Sequence of SCMPDS for s being 0 -started 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,s) = s +* (Start-At (((card I) + 2),SCMPDS)) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started 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,s) = s +* (Start-At (((card I) + 2),SCMPDS)) let s be 0 -started 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,s) = s +* (Start-At (((card I) + 2),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,s) = s +* (Start-At (((card I) + 2),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,s) = s +* (Start-At (((card I) + 2),SCMPDS)) let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) >= 0 implies IExec ((while<0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) ) set d1 = DataLoc ((s . a),i); set WHL = while<0 (a,i,I); set pWHL = stop (while<0 (a,i,I)); set P3 = P +* (stop (while<0 (a,i,I))); set s4 = Comput ((P +* (stop (while<0 (a,i,I)))),s,1); set P4 = P +* (stop (while<0 (a,i,I))); set i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); set SAl = Start-At (((card I) + 2),SCMPDS); I: Initialize s = s by MEMSTR_0:44; A2: IC s = 0 by MEMSTR_0:def_11; A3: while<0 (a,i,I) = ((a,i) >=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; A15: CurInstr ((P +* (stop (while<0 (a,i,I)))),s) = (a,i) >=0_goto ((card I) + 2) by A3, I, SCMPDS_6:11; A6: Comput ((P +* (stop (while<0 (a,i,I)))),s,(0 + 1)) = Following ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,0))) by EXTPRO_1:3 .= Following ((P +* (stop (while<0 (a,i,I)))),s) .= Exec (((a,i) >=0_goto ((card I) + 2)),s) by A3, I, SCMPDS_6:11 ; A7: stop (while<0 (a,i,I)) c= P +* (stop (while<0 (a,i,I))) by FUNCT_4:25; A8: IExec ((while<0 (a,i,I)),P,s) = Result ((P +* (stop (while<0 (a,i,I)))),s) by SCMPDS_4:def_5; assume s . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: IExec ((while<0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) then A9: IC (Comput ((P +* (stop (while<0 (a,i,I)))),s,1)) = ICplusConst (s,((card I) + 2)) by A6, SCMPDS_2:57 .= 0 + ((card I) + 2) by A2, SCMPDS_6:12 ; A12: card (while<0 (a,i,I)) = (card I) + 2 by Th6; then (card I) + 2 in dom (stop (while<0 (a,i,I))) by COMPOS_1:64; then (P +* (stop (while<0 (a,i,I)))) . ((card I) + 2) = (stop (while<0 (a,i,I))) . ((card I) + 2) by A7, GRFUNC_1:2 .= halt SCMPDS by A12, COMPOS_1:64 ; then A13: CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,1))) = halt SCMPDS by A9, PBOOLE:143; then A14: P +* (stop (while<0 (a,i,I))) halts_on s by EXTPRO_1:29; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_0_+_1_holds_ CurInstr_((P_+*_(stop_(while<0_(a,i,I)))),(Comput_((P_+*_(stop_(while<0_(a,i,I)))),s,l)))_<>_halt_SCMPDS let l be Element of NAT ; ::_thesis: ( l < 0 + 1 implies CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,l))) <> halt SCMPDS ) A16: Comput ((P +* (stop (while<0 (a,i,I)))),s,0) = s ; assume l < 0 + 1 ; ::_thesis: CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,l))) <> halt SCMPDS then l = 0 by NAT_1:13; hence CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,l))) <> halt SCMPDS by A15, A16; ::_thesis: verum end; then for l being Element of NAT st CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,l))) = halt SCMPDS holds 1 <= l ; then B17: LifeSpan ((P +* (stop (while<0 (a,i,I)))),s) = 1 by A13, A14, EXTPRO_1:def_15; A19: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((while<0_(a,i,I)),P,s))_holds_ (IExec_((while<0_(a,i,I)),P,s))_._x_=_(s_+*_(Start-At_(((card_I)_+_2),SCMPDS)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((while<0 (a,i,I)),P,s)) implies (IExec ((while<0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 ) A20: dom (Start-At (((card I) + 2),SCMPDS)) = {(IC )} by FUNCOP_1:13; assume A21: x in dom (IExec ((while<0 (a,i,I)),P,s)) ; ::_thesis: (IExec ((while<0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 percases ( x is Int_position or x = IC ) by A21, SCMPDS_4:6; supposeA22: x is Int_position ; ::_thesis: (IExec ((while<0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 then x <> IC by SCMPDS_2:43; then A23: not x in dom (Start-At (((card I) + 2),SCMPDS)) by A20, TARSKI:def_1; thus (IExec ((while<0 (a,i,I)),P,s)) . x = (Comput ((P +* (stop (while<0 (a,i,I)))),s,1)) . x by B17, A8, A14, EXTPRO_1:23 .= s . x by A6, A22, SCMPDS_2:57 .= (s +* (Start-At (((card I) + 2),SCMPDS))) . x by A23, FUNCT_4:11 ; ::_thesis: verum end; supposeA24: x = IC ; ::_thesis: (IExec ((while<0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 hence (IExec ((while<0 (a,i,I)),P,s)) . x = (card I) + 2 by A9, B17, A8, A14, EXTPRO_1:23 .= (s +* (Start-At (((card I) + 2),SCMPDS))) . x by A24, FUNCT_4:113 ; ::_thesis: verum end; end; end; dom (IExec ((while<0 (a,i,I)),P,s)) = the carrier of SCMPDS by PARTFUN1:def_2 .= dom (s +* (Start-At (((card I) + 2),SCMPDS))) by PARTFUN1:def_2 ; hence IExec ((while<0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) by A19, FUNCT_1:2; ::_thesis: verum end; theorem :: SCMPDS_8:11 for P being Instruction-Sequence of SCMPDS for s being 0 -started 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,s)) = (card I) + 2 proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started 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,s)) = (card I) + 2 let s be 0 -started 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,s)) = (card I) + 2 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,s)) = (card I) + 2 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,s)) = (card I) + 2 let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) >= 0 implies IC (IExec ((while<0 (a,i,I)),P,s)) = (card I) + 2 ) assume s . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: IC (IExec ((while<0 (a,i,I)),P,s)) = (card I) + 2 then IExec ((while<0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) by Th10; hence IC (IExec ((while<0 (a,i,I)),P,s)) = (card I) + 2 by FUNCT_4:113; ::_thesis: verum end; theorem :: SCMPDS_8:12 for P being Instruction-Sequence of SCMPDS for s being 0 -started 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,s)) . b = s . b proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started 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,s)) . b = s . b let s be 0 -started 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,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,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,s)) . b = s . b let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) >= 0 implies (IExec ((while<0 (a,i,I)),P,s)) . b = s . b ) assume s . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: (IExec ((while<0 (a,i,I)),P,s)) . b = s . b then A1: IExec ((while<0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) by Th10; not b in dom (Start-At (((card I) + 2),SCMPDS)) by SCMPDS_4:18; hence (IExec ((while<0 (a,i,I)),P,s)) . b = s . b by A1, FUNCT_4:11; ::_thesis: verum end; Lm2: for I being Program of SCMPDS for a being Int_position for i being Integer holds Shift (I,1) 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,1) c= while<0 (a,i,I) let a be Int_position; ::_thesis: for i being Integer holds Shift (I,1) c= while<0 (a,i,I) let i be Integer; ::_thesis: Shift (I,1) c= while<0 (a,i,I) set i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); A1: while<0 (a,i,I) = (((a,i) >=0_goto ((card I) + 2)) ';' I) ';' (Load (goto (- ((card I) + 1)))) by SCMPDS_4:def_3 .= ((Load ((a,i) >=0_goto ((card I) + 2))) ';' I) ';' (Load (goto (- ((card I) + 1)))) by SCMPDS_4:def_2 ; card (Load ((a,i) >=0_goto ((card I) + 2))) = 1 by COMPOS_1:54; hence Shift (I,1) c= while<0 (a,i,I) by A1, SCMPDS_7:3; ::_thesis: verum end; scheme :: SCMPDS_8:sch 1 WhileLHalt{ 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[ State of SCMPDS] } : ( 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 set i1 = (F5(),F6()) >=0_goto ((card F4()) + 2); set i2 = goto (- ((card F4()) + 1)); set WHL = while<0 (F5(),F6(),F4()); set pWHL = 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(t) <= $1 & P1[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] let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1(t) <= k + 1 & P1[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(t) <= k + 1 & P1[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 ) ) T: Initialize t = t by MEMSTR_0:44; assume A7: F1(t) <= k + 1 ; ::_thesis: ( not P1[t] or not t . F5() = F2() . F5() or ( while<0 (F5(),F6(),F4()) is_closed_on t,Q & while<0 (F5(),F6(),F4()) is_halting_on t,Q ) ) assume A8: P1[t] ; ::_thesis: ( not t . F5() = F2() . F5() or ( while<0 (F5(),F6(),F4()) is_closed_on t,Q & while<0 (F5(),F6(),F4()) is_halting_on t,Q ) ) assume A9: 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 ) 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 t,Q & while<0 (F5(),F6(),F4()) is_halting_on t,Q ) hence ( while<0 (F5(),F6(),F4()) is_closed_on t,Q & while<0 (F5(),F6(),F4()) is_halting_on t,Q ) by A9, Th9; ::_thesis: verum end; supposeA10: t . (DataLoc ((F2() . F5()),F6())) < 0 ; ::_thesis: ( while<0 (F5(),F6(),F4()) is_closed_on t,Q & while<0 (F5(),F6(),F4()) is_halting_on t,Q ) A14: 0 in dom (stop (while<0 (F5(),F6(),F4()))) by COMPOS_1:36; A17: while<0 (F5(),F6(),F4()) = ((F5(),F6()) >=0_goto ((card F4()) + 2)) ';' (F4() ';' (goto (- ((card F4()) + 1)))) by SCMPDS_4:15; set Q2 = Q +* (stop F4()); set Q3 = Q +* (stop (while<0 (F5(),F6(),F4()))); set t4 = Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),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())),t); set t5 = Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),t))); set Q5 = Q +* (stop (while<0 (F5(),F6(),F4()))); set l1 = (card F4()) + 1; A21: IC t = 0 by MEMSTR_0:def_11; set m3 = (LifeSpan ((Q +* (stop F4())),t)) + 1; set t6 = Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)); set Q6 = Q +* (stop (while<0 (F5(),F6(),F4()))); set t7 = Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)); set Q7 = Q +* (stop (while<0 (F5(),F6(),F4()))); (card F4()) + 1 < (card F4()) + 2 by XREAL_1:6; then A22: (card F4()) + 1 in dom (while<0 (F5(),F6(),F4())) by Th7; AA: 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 AA, XBOOLE_1:1; Shift (F4(),1) c= while<0 (F5(),F6(),F4()) by Lm2; then A24: Shift (F4(),1) c= Q +* (stop (while<0 (F5(),F6(),F4()))) by A23, XBOOLE_1:1; A25: Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(0 + 1)) = Following ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,0))) by EXTPRO_1:3 .= Following ((Q +* (stop (while<0 (F5(),F6(),F4())))),t) .= Exec (((F5(),F6()) >=0_goto ((card F4()) + 2)),t) by A17, T, SCMPDS_6:11 ; for a being Int_position holds t . a = (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)) . a by A25, SCMPDS_2:57; then A27: DataPart t = DataPart (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)) by SCMPDS_4:8; F4() is_halting_on t,Q by A4, A8, A9, A10; then A28: Q +* (stop F4()) halts_on t by T, SCMPDS_6:def_3; (Q +* (stop F4())) +* (stop F4()) halts_on t by A28; then A30: F4() is_halting_on t,Q +* (stop F4()) by T, SCMPDS_6:def_3; A31: IExec (F4(),Q,t) = Result ((Q +* (stop F4())),t) by SCMPDS_4:def_5; A32: P1[ Initialize (IExec (F4(),Q,t))] by A4, A8, A9, A10; A33: F4() is_closed_on t,Q by A4, A8, A9, A10; A34: F4() is_closed_on t,Q +* (stop F4()) by A4, A8, A9, A10; A35: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)) = succ (IC t) by A10, A25, A9, SCMPDS_2:57 .= 0 + 1 by A21 ; then A36: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),t)))) = (card F4()) + 1 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())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) = (Q +* (stop (while<0 (F5(),F6(),F4())))) . (IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) by PBOOLE:143; A38: Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)) = Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),t))) by EXTPRO_1:4; then A39: CurInstr ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) = (Q +* (stop (while<0 (F5(),F6(),F4())))) . ((card F4()) + 1) by A20, A30, A34, A35, A27, A24, A37, SCMPDS_7:18 .= (while<0 (F5(),F6(),F4())) . ((card F4()) + 1) by A22, A23, GRFUNC_1:2 .= goto (- ((card F4()) + 1)) by Th8 ; A41: Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)) = Following ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card F4()) + 1))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) by A39 ; then IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1))),(0 - ((card F4()) + 1))) by SCMPDS_2:54 .= 0 by A36, A38, SCMPDS_7:1 ; then A42: Initialize (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) = Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)) by MEMSTR_0:46; A43: DataPart (Comput ((Q +* (stop F4())),t,(LifeSpan ((Q +* (stop F4())),t)))) = DataPart (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),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())))),t,1)),(LifeSpan ((Q +* (stop F4())),t)))) = DataPart (Result ((Q +* (stop F4())),t)) by A28, EXTPRO_1:23 .= DataPart (IExec (F4(),Q,t)) by SCMPDS_4:def_5 ; InsCode (goto (- ((card F4()) + 1))) = 14 by SCMPDS_2:12; then InsCode (goto (- ((card F4()) + 1))) in {0,4,5,6,14} by ENUMSET1:def_3; then A45: Initialize (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) = Initialize (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1))) by A41, Th3 .= Initialize (IExec (F4(),Q,t)) by A44, A38, MEMSTR_0:80 ; A46: now__::_thesis:_not_F1((Initialize_(Comput_((Q_+*_(stop_(while<0_(F5(),F6(),F4())))),t,(((LifeSpan_((Q_+*_(stop_F4())),t))_+_1)_+_1)))))_>_k F1((Initialize (IExec (F4(),Q,t)))) < F1(t) by A4, A8, A9, A10; then A47: F1((Initialize (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))))) < k + 1 by A7, A45, XXREAL_0:2; assume F1((Initialize (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),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())))),t,1)),(LifeSpan ((Q +* (stop F4())),t)))) . F5() = (Comput ((Q +* (stop F4())),t,(LifeSpan ((Q +* (stop F4())),t)))) . F5() by A43, SCMPDS_4:8 .= (Result ((Q +* (stop F4())),t)) . F5() by A28, EXTPRO_1:23 .= F2() . F5() by A9, A4, A8, A10, A31 ; A50: (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) . F5() = (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1))) . F5() by A41, SCMPDS_2:54 .= F2() . F5() by A48, EXTPRO_1:4 ; then A51: while<0 (F5(),F6(),F4()) is_closed_on Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)),Q +* (stop (while<0 (F5(),F6(),F4()))) by A6, A32, A45, A46, A42; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(while<0_(F5(),F6(),F4())))),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())))),t,b1)) in dom (stop (while<0 (F5(),F6(),F4()))) percases ( k < ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,b1)) in dom (stop (while<0 (F5(),F6(),F4()))) then A52: k <= (LifeSpan ((Q +* (stop F4())),t)) + 1 by INT_1:7; hereby ::_thesis: verum percases ( k <= LifeSpan ((Q +* (stop F4())),t) or k = (LifeSpan ((Q +* (stop F4())),t)) + 1 ) by A52, NAT_1:8; supposeA53: k <= LifeSpan ((Q +* (stop F4())),t) ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),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())))),t,k)) in dom (stop (while<0 (F5(),F6(),F4()))) hence IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),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())))),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())),t,kn)) as Element of NAT ; kn < k by A54, XREAL_1:29; then kn < LifeSpan ((Q +* (stop F4())),t) by A53, XXREAL_0:2; then (IC (Comput ((Q +* (stop F4())),t,kn))) + 1 = IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,1)),kn)) by A20, A30, A34, A35, A27, A24, SCMPDS_7:16; then A56: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,k)) = lm + 1 by A54, EXTPRO_1:4; IC (Comput ((Q +* (stop F4())),t,kn)) in dom (stop F4()) by A33, T, 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 + 1 <= (card F4()) + 1 by INT_1:7; (card F4()) + 1 < (card F4()) + 3 by XREAL_1:6; then lm + 1 < (card F4()) + 3 by A57, XXREAL_0:2; then lm + 1 < card (stop (while<0 (F5(),F6(),F4()))) by Lm1; hence IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),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())),t)) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,k)) in dom (stop (while<0 (F5(),F6(),F4()))) (card F4()) + 1 in dom (stop (while<0 (F5(),F6(),F4()))) by A22, COMPOS_1:62; hence IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),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())),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,b1)) in dom (stop (while<0 (F5(),F6(),F4()))) then consider nn being Nat such that A59: k = (((LifeSpan ((Q +* (stop F4())),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())))),t,k) = Comput (((Q +* (stop (while<0 (F5(),F6(),F4())))) +* (stop (while<0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))),nn) by A59, EXTPRO_1:4; hence IC (Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,k)) in dom (stop (while<0 (F5(),F6(),F4()))) by A51, A42, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence while<0 (F5(),F6(),F4()) is_closed_on t,Q by T, SCMPDS_6:def_2; ::_thesis: while<0 (F5(),F6(),F4()) is_halting_on t,Q RR: Q +* (stop (while<0 (F5(),F6(),F4()))) = (Q +* (stop (while<0 (F5(),F6(),F4())))) +* (stop (while<0 (F5(),F6(),F4()))) ; while<0 (F5(),F6(),F4()) is_halting_on Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)),Q +* (stop (while<0 (F5(),F6(),F4()))) by A6, A32, A50, A45, A46, A42; then Q +* (stop (while<0 (F5(),F6(),F4()))) halts_on Comput ((Q +* (stop (while<0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)) by A42, RR, SCMPDS_6:def_3; then Q +* (stop (while<0 (F5(),F6(),F4()))) halts_on t by EXTPRO_1:22; hence while<0 (F5(),F6(),F4()) is_halting_on t,Q by T, SCMPDS_6:def_3; ::_thesis: verum end; end; 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(t) <= 0 & P1[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(t) <= 0 & P1[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(t) <= 0 and A62: P1[t] and A63: 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 ) F1(t) = 0 by A61; then t . (DataLoc ((F2() . F5()),F6())) >= 0 by A2, A62; hence ( while<0 (F5(),F6(),F4()) is_closed_on t,Q & while<0 (F5(),F6(),F4()) is_halting_on t,Q ) by A63, Th9; ::_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; ::_thesis: verum end; scheme :: SCMPDS_8:sch 2 WhileLExec{ 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[ State of SCMPDS] } : 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 II: Initialize F2() = F2() by MEMSTR_0:44; set WHL = while<0 (F5(),F6(),F4()); set pWHL = stop (while<0 (F5(),F6(),F4())); set P1 = F3() +* (stop (while<0 (F5(),F6(),F4()))); set PI = F3() +* (stop F4()); set m1 = (LifeSpan ((F3() +* (stop F4())),F2())) + 2; set s2 = Initialize (IExec (F4(),F3(),F2())); set P2 = F3() +* (stop (while<0 (F5(),F6(),F4()))); set m2 = LifeSpan ((F3() +* (stop (while<0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))); A9: P1[F2()] by A4; A10: stop F4() c= F3() +* (stop F4()) by FUNCT_4:25; A11: F4() is_closed_on F2(),F3() +* (stop F4()) by A2, A4, A5; F4() is_halting_on F2(),F3() by A2, A4, A5; then A12: F3() +* (stop F4()) halts_on F2() by II, SCMPDS_6:def_3; (F3() +* (stop F4())) +* (stop F4()) halts_on F2() by A12; then A14: F4() is_halting_on F2(),F3() +* (stop F4()) by II, SCMPDS_6:def_3; set Es = IExec (F4(),F3(),F2()); set bj = DataLoc (((Initialize (IExec (F4(),F3(),F2()))) . F5()),F6()); set EP = F3(); A15: 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))] ) by A5; A16: for t being 0 -started State of SCMPDS st P1[t] & F1(t) = 0 holds t . (DataLoc ((F2() . F5()),F6())) >= 0 by A3; ( while<0 (F5(),F6(),F4()) is_closed_on F2(),F3() & while<0 (F5(),F6(),F4()) is_halting_on F2(),F3() ) from SCMPDS_8:sch_1(A16, A9, A15); then A17: F3() +* (stop (while<0 (F5(),F6(),F4()))) halts_on F2() by II, SCMPDS_6:def_3; B18: (IExec (F4(),F3(),F2())) . F5() = (Initialize (IExec (F4(),F3(),F2()))) . F5() by SCMPDS_5:15; A18: (IExec (F4(),F3(),F2())) . F5() = F2() . F5() by A2, A4, A5; then A19: for t being 0 -started State of SCMPDS st P1[t] & F1(t) = 0 holds t . (DataLoc (((Initialize (IExec (F4(),F3(),F2()))) . F5()),F6())) >= 0 by A3, B18; A20: 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 & F1((Initialize (IExec (F4(),Q,t)))) < F1(t) & P1[ Initialize (IExec (F4(),Q,t))] ) by A5, A18, B18; A21: P1[ Initialize (IExec (F4(),F3(),F2()))] by A2, A4, A5; ( 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 SCMPDS_8:sch_1(A19, A21, A20); then A22: F3() +* (stop (while<0 (F5(),F6(),F4()))) halts_on Initialize (Initialize (IExec (F4(),F3(),F2()))) 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 ((card F4()) + 2); set i2 = goto (- ((card F4()) + 1)); set b = DataLoc ((F2() . F5()),F6()); A23: while<0 (F5(),F6(),F4()) = ((F5(),F6()) >=0_goto ((card F4()) + 2)) ';' (F4() ';' (goto (- ((card F4()) + 1)))) by SCMPDS_4:15; 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 l1 = (card F4()) + 1; A24: IC F2() = 0 by MEMSTR_0:def_11; A25: 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 ((card F4()) + 2)),F2()) by A23, II, SCMPDS_6:11 ; for a being Int_position holds F2() . a = (Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),1)) . a by A25, SCMPDS_2:57; then A27: DataPart F2() = DataPart (Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),1)) by SCMPDS_4:8; 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()) + 1 < (card F4()) + 2 by XREAL_1:6; then A28: (card F4()) + 1 in dom (while<0 (F5(),F6(),F4())) by Th7; set m0 = LifeSpan ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2()); set s7 = Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1)); XX: while<0 (F5(),F6(),F4()) c= stop (while<0 (F5(),F6(),F4())) by AFINSQ_1:74; stop (while<0 (F5(),F6(),F4())) c= F3() +* (stop (while<0 (F5(),F6(),F4()))) by FUNCT_4:25; then A33: while<0 (F5(),F6(),F4()) c= F3() +* (stop (while<0 (F5(),F6(),F4()))) by XX, XBOOLE_1:1; Shift (F4(),1) c= while<0 (F5(),F6(),F4()) by Lm2; then A34: Shift (F4(),1) c= F3() +* (stop (while<0 (F5(),F6(),F4()))) by A33, XBOOLE_1:1; A35: IC (Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),1)) = succ (IC F2()) by A2, A25, SCMPDS_2:57 .= 0 + 1 by A24 ; then A36: 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()) + 1 by A10, A14, A11, A27, A34, SCMPDS_7:18; A37: (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; A38: 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 A39: 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()) + 1) by A10, A14, A11, A35, A27, A34, A37, SCMPDS_7:18 .= (while<0 (F5(),F6(),F4())) . ((card F4()) + 1) by A28, A33, GRFUNC_1:2 .= goto (- ((card F4()) + 1)) by Th8 ; A41: 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()) + 1))),(Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) by A39 ; 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()) + 1))) by SCMPDS_2:54 .= 0 by A36, A38, SCMPDS_7:1 ; then A42: 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:def_11; A43: 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 A10, A14, A11, A35, A27, A34, 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 A44: 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 A43, SCMPDS_4:8 .= (Result ((F3() +* (stop F4())),F2())) . x by A12, EXTPRO_1:23 .= (IExec (F4(),F3(),F2())) . x by SCMPDS_4:def_5 ; hence (Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1))) . x = (IExec (F4(),F3(),F2())) . x by A38, A41, SCMPDS_2:54 .= (Initialize (IExec (F4(),F3(),F2()))) . x by A44, FUNCT_4:11 ; ::_thesis: verum end; then A46: 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; A48: Comput ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 2)) = Initialize (IExec (F4(),F3(),F2())) by A46, A42, 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 ((card F4()) + 2) by A23, SCMPDS_6:11; then LifeSpan ((F3() +* (stop (while<0 (F5(),F6(),F4())))),F2()) > (LifeSpan ((F3() +* (stop F4())),F2())) + 2 by A17, EXTPRO_1:36, SCMPDS_6:18; 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; 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()))))))) = 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()))))))))) = halt SCMPDS by A22, 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 A17, 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 A17, EXTPRO_1:def_15; then nn >= LifeSpan ((F3() +* (stop (while<0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))) by A22, 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 A17, 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 A22, EXTPRO_1:23 .= IExec ((while<0 (F5(),F6(),F4())),F3(),(Initialize (IExec (F4(),F3(),F2())))) by SCMPDS_4:def_5 ; ::_thesis: verum end; theorem :: SCMPDS_8:13 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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: for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let X be set ; ::_thesis: for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let f be Function of (product (the_Values_of SCMPDS)),NAT; ::_thesis: ( ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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); set WHL = while<0 (a,i,I); set pWHL = stop (while<0 (a,i,I)); set pI = stop I; set i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st f . t <= $1 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a holds ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ); assume A2: for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 & not ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) or ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) ) assume A3: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & f . (Initialize (IExec (I,Q,t))) < f . t & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_thesis: ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) A4: for k being Element of NAT st S1[k] holds S1[k + 1] proof let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A5: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_f_._t_<=_k_+_1_&_(_for_x_being_Int_position_st_x_in_X_holds_ t_._x_=_s_._x_)_&_t_._a_=_s_._a_holds_ (_while<0_(a,i,I)_is_closed_on_t,Q_&_while<0_(a,i,I)_is_halting_on_t,Q_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st f . t <= k + 1 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a holds ( while<0 (a,i,I) is_closed_on b2,b3 & while<0 (a,i,I) is_halting_on b2,b3 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( f . t <= k + 1 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a implies ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) ) T: Initialize t = t by MEMSTR_0:44; assume A6: f . t <= k + 1 ; ::_thesis: ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a implies ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) ) assume A7: for x being Int_position st x in X holds t . x = s . x ; ::_thesis: ( t . a = s . a implies ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) ) assume A8: t . a = s . a ; ::_thesis: ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) percases ( t . (DataLoc ((s . a),i)) >= 0 or t . (DataLoc ((s . a),i)) < 0 ) ; suppose t . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) hence ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) by A8, Th9; ::_thesis: verum end; supposeA9: t . (DataLoc ((s . a),i)) < 0 ; ::_thesis: ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) A13: 0 in dom (stop (while<0 (a,i,I))) by COMPOS_1:36; A15: while<0 (a,i,I) = ((a,i) >=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; A16: f . (Initialize (IExec (I,Q,t))) < f . t by A3, A7, A8, A9; set t2 = t; set Q2 = Q +* (stop I); set t3 = t; set Q3 = Q +* (stop (while<0 (a,i,I))); set t4 = Comput ((Q +* (stop (while<0 (a,i,I)))),t,1); set Q4 = Q +* (stop (while<0 (a,i,I))); A20: stop I c= Q +* (stop I) by FUNCT_4:25; A21: Comput ((Q +* (stop (while<0 (a,i,I)))),t,(0 + 1)) = Following ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,0))) by EXTPRO_1:3 .= Following ((Q +* (stop (while<0 (a,i,I)))),t) .= Exec (((a,i) >=0_goto ((card I) + 2)),t) by A15, T, SCMPDS_6:11 ; for a being Int_position holds t . a = (Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)) . a by A21, SCMPDS_2:57; then A23: DataPart t = DataPart (Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)) by SCMPDS_4:8; XX: while<0 (a,i,I) c= stop (while<0 (a,i,I)) by AFINSQ_1:74; stop (while<0 (a,i,I)) c= Q +* (stop (while<0 (a,i,I))) by FUNCT_4:25; then A24: while<0 (a,i,I) c= Q +* (stop (while<0 (a,i,I))) by XX, XBOOLE_1:1; Shift (I,1) c= while<0 (a,i,I) by Lm2; then A25: Shift (I,1) c= Q +* (stop (while<0 (a,i,I))) by A24, XBOOLE_1:1; set m2 = LifeSpan ((Q +* (stop I)),t); set t5 = Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t))); set Q5 = Q +* (stop (while<0 (a,i,I))); set l1 = (card I) + 1; A26: IC t = 0 by MEMSTR_0:def_11; set m3 = (LifeSpan ((Q +* (stop I)),t)) + 1; set t6 = Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)); set Q6 = Q +* (stop (while<0 (a,i,I))); set t7 = Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)); set Q7 = Q +* (stop (while<0 (a,i,I))); (card I) + 1 < (card I) + 2 by XREAL_1:6; then A27: (card I) + 1 in dom (while<0 (a,i,I)) by Th7; A28: I is_closed_on t,Q by A3, A7, A8, A9; A29: I is_closed_on t,Q +* (stop I) by A3, A7, A8, A9; I is_halting_on t,Q by A3, A7, A8, A9; then A30: Q +* (stop I) halts_on t by T, SCMPDS_6:def_3; (Q +* (stop I)) +* (stop I) halts_on t by A30; then A32: I is_halting_on t,Q +* (stop I) by T, SCMPDS_6:def_3; A33: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)) = succ (IC t) by A9, A21, A8, SCMPDS_2:57 .= 0 + 1 by A26 ; then A34: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) = (card I) + 1 by A20, A32, A29, A23, A25, SCMPDS_7:18; A35: (Q +* (stop (while<0 (a,i,I)))) /. (IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) = (Q +* (stop (while<0 (a,i,I)))) . (IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) by PBOOLE:143; A36: Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)) = Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t))) by EXTPRO_1:4; then A37: CurInstr ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) = (Q +* (stop (while<0 (a,i,I)))) . ((card I) + 1) by A20, A32, A29, A33, A23, A25, A35, SCMPDS_7:18 .= (while<0 (a,i,I)) . ((card I) + 1) by A27, A24, GRFUNC_1:2 .= goto (- ((card I) + 1)) by Th8 ; A38: DataPart (Comput ((Q +* (stop I)),t,(LifeSpan ((Q +* (stop I)),t)))) = DataPart (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) by A20, A32, A29, A33, A23, A25, SCMPDS_7:18; then A39: DataPart (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) = DataPart (Result ((Q +* (stop I)),t)) by A30, EXTPRO_1:23 .= DataPart (IExec (I,Q,t)) by SCMPDS_4:def_5 ; A41: Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)) = Following ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card I) + 1))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) by A37 ; then IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1))),(0 - ((card I) + 1))) by SCMPDS_2:54 .= 0 by A34, A36, SCMPDS_7:1 ; then A42: Initialize (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) = Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)) by MEMSTR_0:46; A43: IExec (I,Q,t) = Result ((Q +* (stop I)),t) by SCMPDS_4:def_5; A44: now__::_thesis:_for_x_being_Int_position_st_x_in_X_holds_ (Comput_((Q_+*_(stop_(while<0_(a,i,I)))),t,(((LifeSpan_((Q_+*_(stop_I)),t))_+_1)_+_1)))_._x_=_s_._x let x be Int_position; ::_thesis: ( x in X implies (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = s . x ) assume A45: x in X ; ::_thesis: (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = s . x (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) . x = (Comput ((Q +* (stop I)),t,(LifeSpan ((Q +* (stop I)),t)))) . x by A38, SCMPDS_4:8 .= (Result ((Q +* (stop I)),t)) . x by A30, EXTPRO_1:23 .= (IExec (I,Q,t)) . x by SCMPDS_4:def_5 .= t . x by A3, A7, A8, A9, A45 .= s . x by A7, A45 ; hence (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = s . x by A36, A41, SCMPDS_2:54; ::_thesis: verum end; InsCode (goto (- ((card I) + 1))) = 14 by SCMPDS_2:12; then InsCode (goto (- ((card I) + 1))) in {0,4,5,6,14} by ENUMSET1:def_3; then A47: Initialize (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) = Initialize (Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1))) by A41, Th3 .= Initialize (IExec (I,Q,t)) by A39, A36, MEMSTR_0:80 ; A48: now__::_thesis:_not_f_._(Comput_((Q_+*_(stop_(while<0_(a,i,I)))),t,(((LifeSpan_((Q_+*_(stop_I)),t))_+_1)_+_1)))_>_k assume A49: f . (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) > k ; ::_thesis: contradiction f . (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) < k + 1 by A6, A16, A47, A42, XXREAL_0:2; hence contradiction by A49, INT_1:7; ::_thesis: verum end; A50: (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) . a = (Comput ((Q +* (stop I)),t,(LifeSpan ((Q +* (stop I)),t)))) . a by A38, SCMPDS_4:8 .= (Result ((Q +* (stop I)),t)) . a by A30, EXTPRO_1:23 .= s . a by A8, A3, A7, A9, A43 ; A52: (Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . a = (Comput ((Q +* (stop (while<0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1))) . a by A41, SCMPDS_2:54 .= s . a by A50, EXTPRO_1:4 ; then A53: while<0 (a,i,I) is_closed_on Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)),Q +* (stop (while<0 (a,i,I))) by A5, A44, A48, A42; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(while<0_(a,i,I)))),t,k))_in_dom_(stop_(while<0_(a,i,I))) let k be Element of NAT ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,b1)) in dom (stop (while<0 (a,i,I))) percases ( k < ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,b1)) in dom (stop (while<0 (a,i,I))) then A54: k <= (LifeSpan ((Q +* (stop I)),t)) + 1 by INT_1:7; hereby ::_thesis: verum percases ( k <= LifeSpan ((Q +* (stop I)),t) or k = (LifeSpan ((Q +* (stop I)),t)) + 1 ) by A54, NAT_1:8; supposeA55: k <= LifeSpan ((Q +* (stop I)),t) ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) hereby ::_thesis: verum percases ( k = 0 or k <> 0 ) ; suppose k = 0 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) by A13, A26, EXTPRO_1:2; ::_thesis: verum end; suppose k <> 0 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) then consider kn being Nat such that A56: k = kn + 1 by NAT_1:6; reconsider kn = kn as Element of NAT by ORDINAL1:def_12; reconsider lm = IC (Comput ((Q +* (stop I)),t,kn)) as Element of NAT ; kn < k by A56, XREAL_1:29; then kn < LifeSpan ((Q +* (stop I)),t) by A55, XXREAL_0:2; then (IC (Comput ((Q +* (stop I)),t,kn))) + 1 = IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,1)),kn)) by A20, A32, A29, A33, A23, A25, SCMPDS_7:16; then A58: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) = lm + 1 by A56, EXTPRO_1:4; IC (Comput ((Q +* (stop I)),t,kn)) in dom (stop I) by A28, T, SCMPDS_6:def_2; then lm < card (stop I) by AFINSQ_1:66; then lm < (card I) + 1 by COMPOS_1:55; then A59: lm + 1 <= (card I) + 1 by INT_1:7; (card I) + 1 < (card I) + 3 by XREAL_1:6; then lm + 1 < (card I) + 3 by A59, XXREAL_0:2; then lm + 1 < card (stop (while<0 (a,i,I))) by Lm1; hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) by A58, AFINSQ_1:66; ::_thesis: verum end; end; end; end; supposeA60: k = (LifeSpan ((Q +* (stop I)),t)) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) (card I) + 1 in dom (stop (while<0 (a,i,I))) by A27, COMPOS_1:62; hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) by A20, A32, A29, A33, A23, A25, A36, A60, SCMPDS_7:18; ::_thesis: verum end; end; end; end; suppose k >= ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,b1)) in dom (stop (while<0 (a,i,I))) then consider nn being Nat such that A61: k = (((LifeSpan ((Q +* (stop I)),t)) + 1) + 1) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; Comput ((Q +* (stop (while<0 (a,i,I)))),t,k) = Comput (((Q +* (stop (while<0 (a,i,I)))) +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))),nn) by A61, EXTPRO_1:4; hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),t,k)) in dom (stop (while<0 (a,i,I))) by A53, A42, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence while<0 (a,i,I) is_closed_on t,Q by T, SCMPDS_6:def_2; ::_thesis: while<0 (a,i,I) is_halting_on t,Q RR: Q +* (stop (while<0 (a,i,I))) = (Q +* (stop (while<0 (a,i,I)))) +* (stop (while<0 (a,i,I))) ; while<0 (a,i,I) is_halting_on Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)),Q +* (stop (while<0 (a,i,I))) by A5, A52, A44, A48, A42; then Q +* (stop (while<0 (a,i,I))) halts_on Comput ((Q +* (stop (while<0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)) by A42, RR, SCMPDS_6:def_3; then Q +* (stop (while<0 (a,i,I))) halts_on t by EXTPRO_1:22; hence while<0 (a,i,I) is_halting_on t,Q by T, SCMPDS_6:def_3; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; set n = f . s; A62: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st f . t <= 0 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a holds ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( f . t <= 0 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a implies ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) ) assume f . t <= 0 ; ::_thesis: ( ex x being Int_position st ( x in X & not t . x = s . x ) or not t . a = s . a or ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) ) then f . t = 0 ; then A63: t . (DataLoc ((s . a),i)) >= 0 by A2; assume for x being Int_position st x in X holds t . x = s . x ; ::_thesis: ( not t . a = s . a or ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) ) assume t . a = s . a ; ::_thesis: ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) hence ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) by A63, Th9; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A62, A4); then A64: S1[f . s] ; for x being Int_position st x in X holds s . x = s . x ; hence ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) by A64; ::_thesis: verum end; theorem :: SCMPDS_8:14 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let a be Int_position; ::_thesis: for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let i be Integer; ::_thesis: for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let X be set ; ::_thesis: for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let f be Function of (product (the_Values_of SCMPDS)),NAT; ::_thesis: ( s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) implies IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) set b = DataLoc ((s . a),i); deffunc H1( State of SCMPDS) -> Element of NAT = f . $1; defpred S1[ State of SCMPDS] means for x being Int_position st x in X holds $1 . x = s . x; assume A2: s . (DataLoc ((s . a),i)) < 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS st ( f . t = 0 & not t . (DataLoc ((s . a),i)) >= 0 ) or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) or IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) assume for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) or IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) then A3: for t being 0 -started State of SCMPDS st S1[t] & H1(t) = 0 holds t . (DataLoc ((s . a),i)) >= 0 ; assume A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_thesis: IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) A5: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_S1[t]_&_t_._a_=_s_._a_&_t_._(DataLoc_((s_._a),i))_<_0_holds_ (_(IExec_(I,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))] ) ) set v = t; assume that A6: S1[t] and A7: ( t . a = s . a & 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); thus ( (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) ) by A4, A7, A6; ::_thesis: S1[ Initialize (IExec (I,Q,t))] thus S1[ Initialize (IExec (I,Q,t))] ::_thesis: verum proof set v = Initialize (IExec (I,Q,t)); let x be Int_position; ::_thesis: ( x in X implies (Initialize (IExec (I,Q,t))) . x = s . x ) assume A9: x in X ; ::_thesis: (Initialize (IExec (I,Q,t))) . x = s . x then (IExec (I,Q,t)) . x = t . x by A4, A7, A6; then (Initialize (IExec (I,Q,t))) . x = t . x by SCMPDS_5:15; hence (Initialize (IExec (I,Q,t))) . x = s . x by A6, A9; ::_thesis: verum end; end; A10: S1[s] ; IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) from SCMPDS_8:sch_2(A2, A3, A10, A5); hence IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ; ::_thesis: verum end; theorem Th15: :: SCMPDS_8:15 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( 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 being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( 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 being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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: for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) let X be set ; ::_thesis: ( ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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); set WHL = while<0 (a,i,I); set pWHL = stop (while<0 (a,i,I)); set pI = stop I; set i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st - (t . (DataLoc ((s . a),i))) <= $1 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a holds ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ); assume A2: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_thesis: ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) A3: 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 A4: S1[k] ; ::_thesis: S1[k + 1] now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_-_(t_._(DataLoc_((s_._a),i)))_<=_k_+_1_&_(_for_x_being_Int_position_st_x_in_X_holds_ t_._x_=_s_._x_)_&_t_._a_=_s_._a_holds_ (_while<0_(a,i,I)_is_closed_on_t,Q_&_while<0_(a,i,I)_is_halting_on_t,Q_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st - (t . (DataLoc ((s . a),i))) <= k + 1 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a holds ( while<0 (a,i,I) is_closed_on b2,b3 & while<0 (a,i,I) is_halting_on b2,b3 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( - (t . (DataLoc ((s . a),i))) <= k + 1 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a implies ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) ) T: Initialize t = t by MEMSTR_0:44; assume A5: - (t . (DataLoc ((s . a),i))) <= k + 1 ; ::_thesis: ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a implies ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) ) assume A6: for x being Int_position st x in X holds t . x = s . x ; ::_thesis: ( t . a = s . a implies ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) ) assume A7: t . a = s . a ; ::_thesis: ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) percases ( t . (DataLoc ((s . a),i)) >= 0 or t . (DataLoc ((s . a),i)) < 0 ) ; suppose t . (DataLoc ((s . a),i)) >= 0 ; ::_thesis: ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) hence ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) by A7, Th9; ::_thesis: verum end; supposeA8: t . (DataLoc ((s . a),i)) < 0 ; ::_thesis: ( while<0 (a,i,I) is_closed_on b1,b2 & while<0 (a,i,I) is_halting_on b1,b2 ) A12: (IExec (I,Q,(Initialize t))) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) by A2, A6, A7, A8, T; A13: 0 in dom (stop (while<0 (a,i,I))) by COMPOS_1:36; A15: not DataLoc ((s . a),i) in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; A16: while<0 (a,i,I) = ((a,i) >=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; set t2 = Initialize t; set Q2 = Q +* (stop I); set t3 = Initialize t; set Q3 = Q +* (stop (while<0 (a,i,I))); set t4 = Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1); set Q4 = Q +* (stop (while<0 (a,i,I))); A19: stop I c= Q +* (stop I) by FUNCT_4:25; A20: Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(0 + 1)) = Following ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),0))) by EXTPRO_1:3 .= Following ((Q +* (stop (while<0 (a,i,I)))),(Initialize t)) .= Exec (((a,i) >=0_goto ((card I) + 2)),(Initialize t)) by A16, SCMPDS_6:11 ; for a being Int_position holds (Initialize t) . a = (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)) . a by A20, SCMPDS_2:57; then A22: DataPart (Initialize t) = DataPart (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)) by SCMPDS_4:8; XX: while<0 (a,i,I) c= stop (while<0 (a,i,I)) by AFINSQ_1:74; stop (while<0 (a,i,I)) c= Q +* (stop (while<0 (a,i,I))) by FUNCT_4:25; then A23: while<0 (a,i,I) c= Q +* (stop (while<0 (a,i,I))) by XX, XBOOLE_1:1; Shift (I,1) c= while<0 (a,i,I) by Lm2; then A24: Shift (I,1) c= Q +* (stop (while<0 (a,i,I))) by A23, XBOOLE_1:1; set m2 = LifeSpan ((Q +* (stop I)),(Initialize t)); set t5 = Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t)))); set Q5 = Q +* (stop (while<0 (a,i,I))); set l1 = (card I) + 1; A25: IC (Initialize t) = 0 by MEMSTR_0:def_11; set m3 = (LifeSpan ((Q +* (stop I)),(Initialize t))) + 1; set t6 = Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)); set Q6 = Q +* (stop (while<0 (a,i,I))); set t7 = Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)); set Q7 = Q +* (stop (while<0 (a,i,I))); (card I) + 1 < (card I) + 2 by XREAL_1:6; then A26: (card I) + 1 in dom (while<0 (a,i,I)) by Th7; A27: IExec (I,Q,(Initialize t)) = Result ((Q +* (stop I)),(Initialize t)) by SCMPDS_4:def_5; A28: I is_closed_on t,Q by A2, A6, A7, A8; then A29: I is_closed_on Initialize t,Q +* (stop I) by SCMPDS_6:24; I is_halting_on t,Q by A2, A6, A7, A8; then A30: Q +* (stop I) halts_on Initialize t by SCMPDS_6:def_3; (Q +* (stop I)) +* (stop I) halts_on Initialize (Initialize t) by A30; then A32: I is_halting_on Initialize t,Q +* (stop I) by SCMPDS_6:def_3; not a in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; then (Initialize t) . (DataLoc (((Initialize t) . a),i)) = (Initialize t) . (DataLoc ((s . a),i)) by A7, FUNCT_4:11 .= t . (DataLoc ((s . a),i)) by A15, FUNCT_4:11 ; then A33: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)) = succ (IC (Initialize t)) by A8, A20, SCMPDS_2:57 .= 0 + 1 by A25 ; then A34: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t))))) = (card I) + 1 by A19, A32, A29, A22, A24, SCMPDS_7:18; A35: (Q +* (stop (while<0 (a,i,I)))) /. (IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)))) = (Q +* (stop (while<0 (a,i,I)))) . (IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)))) by PBOOLE:143; A36: Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)) = Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t)))) by EXTPRO_1:4; then A37: CurInstr ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)))) = (Q +* (stop (while<0 (a,i,I)))) . ((card I) + 1) by A19, A32, A29, A33, A22, A24, A35, SCMPDS_7:18 .= (while<0 (a,i,I)) . ((card I) + 1) by A26, A23, GRFUNC_1:2 .= goto (- ((card I) + 1)) by Th8 ; A38: DataPart (Comput ((Q +* (stop I)),(Initialize t),(LifeSpan ((Q +* (stop I)),(Initialize t))))) = DataPart (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t))))) by A19, A32, A29, A33, A22, A24, SCMPDS_7:18; then A39: (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t))))) . a = (Comput ((Q +* (stop I)),(Initialize t),(LifeSpan ((Q +* (stop I)),(Initialize t))))) . a by SCMPDS_4:8 .= (Result ((Q +* (stop I)),(Initialize t))) . a by A30, EXTPRO_1:23 .= s . a by A7, A2, A6, A8, A27, T ; A41: Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)) = Following ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card I) + 1))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1)))) by A37 ; then IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1))),(0 - ((card I) + 1))) by SCMPDS_2:54 .= 0 by A34, A36, SCMPDS_7:1 ; then A42: Initialize (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) = Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)) by MEMSTR_0:46; A43: now__::_thesis:_for_x_being_Int_position_st_x_in_X_holds_ (Comput_((Q_+*_(stop_(while<0_(a,i,I)))),(Initialize_t),(((LifeSpan_((Q_+*_(stop_I)),(Initialize_t)))_+_1)_+_1)))_._x_=_s_._x let x be Int_position; ::_thesis: ( x in X implies (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . x = s . x ) assume A44: x in X ; ::_thesis: (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . x = s . x (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t))))) . x = (Comput ((Q +* (stop I)),(Initialize t),(LifeSpan ((Q +* (stop I)),(Initialize t))))) . x by A38, SCMPDS_4:8 .= (Result ((Q +* (stop I)),(Initialize t))) . x by A30, EXTPRO_1:23 .= (IExec (I,Q,(Initialize t))) . x by SCMPDS_4:def_5 .= t . x by A2, A6, A7, A8, A44, T .= s . x by A6, A44 ; hence (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . x = s . x by A36, A41, SCMPDS_2:54; ::_thesis: verum end; (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),(LifeSpan ((Q +* (stop I)),(Initialize t))))) . (DataLoc ((s . a),i)) = (Comput ((Q +* (stop I)),(Initialize t),(LifeSpan ((Q +* (stop I)),(Initialize t))))) . (DataLoc ((s . a),i)) by A38, SCMPDS_4:8 .= (Result ((Q +* (stop I)),(Initialize t))) . (DataLoc ((s . a),i)) by A30, EXTPRO_1:23 .= (IExec (I,Q,(Initialize t))) . (DataLoc ((s . a),i)) by SCMPDS_4:def_5 ; then A46: (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . (DataLoc ((s . a),i)) = (IExec (I,Q,(Initialize t))) . (DataLoc ((s . a),i)) by A36, A41, SCMPDS_2:54; A47: now__::_thesis:_not_-_((Comput_((Q_+*_(stop_(while<0_(a,i,I)))),(Initialize_t),(((LifeSpan_((Q_+*_(stop_I)),(Initialize_t)))_+_1)_+_1)))_._(DataLoc_((s_._a),i)))_>_k - ((Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . (DataLoc ((s . a),i))) < - (t . (DataLoc ((s . a),i))) by A12, A46, XREAL_1:24; then A48: - ((Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . (DataLoc ((s . a),i))) < k + 1 by A5, XXREAL_0:2; assume - ((Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . (DataLoc ((s . a),i))) > k ; ::_thesis: contradiction hence contradiction by A48, INT_1:7; ::_thesis: verum end; A50: (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1))) . a = (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1))) . a by A41, SCMPDS_2:54 .= s . a by A39, EXTPRO_1:4 ; then A51: while<0 (a,i,I) is_closed_on Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)),Q +* (stop (while<0 (a,i,I))) by A4, A43, A47, A42; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(while<0_(a,i,I)))),(Initialize_t),k))_in_dom_(stop_(while<0_(a,i,I))) let k be Element of NAT ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),b1)) in dom (stop (while<0 (a,i,I))) percases ( k < ((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),b1)) in dom (stop (while<0 (a,i,I))) then A52: k <= (LifeSpan ((Q +* (stop I)),(Initialize t))) + 1 by INT_1:7; hereby ::_thesis: verum percases ( k <= LifeSpan ((Q +* (stop I)),(Initialize t)) or k = (LifeSpan ((Q +* (stop I)),(Initialize t))) + 1 ) by A52, NAT_1:8; supposeA53: k <= LifeSpan ((Q +* (stop I)),(Initialize t)) ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) hereby ::_thesis: verum percases ( k = 0 or k <> 0 ) ; suppose k = 0 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) by A13, A25, EXTPRO_1:2; ::_thesis: verum end; suppose k <> 0 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) 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 I)),(Initialize t),kn)) as Element of NAT ; kn < k by A54, XREAL_1:29; then kn < LifeSpan ((Q +* (stop I)),(Initialize t)) by A53, XXREAL_0:2; then (IC (Comput ((Q +* (stop I)),(Initialize t),kn))) + 1 = IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),1)),kn)) by A19, A32, A29, A33, A22, A24, SCMPDS_7:16; then A56: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) = lm + 1 by A54, EXTPRO_1:4; IC (Comput ((Q +* (stop I)),(Initialize t),kn)) in dom (stop I) by A28, SCMPDS_6:def_2; then lm < card (stop I) by AFINSQ_1:66; then lm < (card I) + 1 by COMPOS_1:55; then A57: lm + 1 <= (card I) + 1 by INT_1:7; (card I) + 1 < (card I) + 3 by XREAL_1:6; then lm + 1 < (card I) + 3 by A57, XXREAL_0:2; then lm + 1 < card (stop (while<0 (a,i,I))) by Lm1; hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) by A56, AFINSQ_1:66; ::_thesis: verum end; end; end; end; supposeA58: k = (LifeSpan ((Q +* (stop I)),(Initialize t))) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) (card I) + 1 in dom (stop (while<0 (a,i,I))) by A26, COMPOS_1:62; hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) by A19, A32, A29, A33, A22, A24, A36, A58, SCMPDS_7:18; ::_thesis: verum end; end; end; end; suppose k >= ((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),b1)) in dom (stop (while<0 (a,i,I))) then consider nn being Nat such that A59: k = (((LifeSpan ((Q +* (stop I)),(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 (a,i,I)))),(Initialize t),k) = Comput (((Q +* (stop (while<0 (a,i,I)))) +* (stop (while<0 (a,i,I)))),(Initialize (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)))),nn) by A42, A59, EXTPRO_1:4; hence IC (Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),k)) in dom (stop (while<0 (a,i,I))) by A51, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence while<0 (a,i,I) is_closed_on t,Q by SCMPDS_6:def_2; ::_thesis: while<0 (a,i,I) is_halting_on t,Q RR: Q +* (stop (while<0 (a,i,I))) = (Q +* (stop (while<0 (a,i,I)))) +* (stop (while<0 (a,i,I))) ; while<0 (a,i,I) is_halting_on Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)),Q +* (stop (while<0 (a,i,I))) by A4, A50, A43, A47, A42; then Q +* (stop (while<0 (a,i,I))) halts_on Comput ((Q +* (stop (while<0 (a,i,I)))),(Initialize t),(((LifeSpan ((Q +* (stop I)),(Initialize t))) + 1) + 1)) by A42, RR, SCMPDS_6:def_3; then Q +* (stop (while<0 (a,i,I))) halts_on Initialize t by EXTPRO_1:22; hence while<0 (a,i,I) is_halting_on t,Q by SCMPDS_6:def_3; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; A60: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st - (t . (DataLoc ((s . a),i))) <= 0 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a holds ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( - (t . (DataLoc ((s . a),i))) <= 0 & ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a implies ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) ) assume - (t . (DataLoc ((s . a),i))) <= 0 ; ::_thesis: ( ex x being Int_position st ( x in X & not t . x = s . x ) or not t . a = s . a or ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) ) then - (t . (DataLoc ((s . a),i))) <= - 0 ; then A61: t . (DataLoc ((s . a),i)) >= 0 by XREAL_1:24; assume for x being Int_position st x in X holds t . x = s . x ; ::_thesis: ( not t . a = s . a or ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) ) assume t . a = s . a ; ::_thesis: ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) hence ( while<0 (a,i,I) is_closed_on t,Q & while<0 (a,i,I) is_halting_on t,Q ) by A61, Th9; ::_thesis: verum end; A62: for k being Element of NAT holds S1[k] from NAT_1:sch_1(A60, A3); percases ( s . (DataLoc ((s . a),i)) >= 0 or s . (DataLoc ((s . a),i)) < 0 ) ; suppose 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 ) hence ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) by Th9; ::_thesis: verum end; suppose 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 ) then reconsider n = - (s . (DataLoc ((s . a),i))) as Element of NAT by INT_1:3; ( S1[n] & ( for x being Int_position st x in X holds s . x = s . x ) ) by A62; hence ( while<0 (a,i,I) is_closed_on s,P & while<0 (a,i,I) is_halting_on s,P ) ; ::_thesis: verum end; end; end; theorem :: SCMPDS_8: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 being Int_position for i being Integer for X being set st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i being Integer for X being set st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i being Integer for X being set st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i being Integer for X being set st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let a be Int_position; ::_thesis: for i being Integer for X being set st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let i be Integer; ::_thesis: for X being set st s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let X be set ; ::_thesis: ( s . (DataLoc ((s . a),i)) < 0 & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) implies IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) I: Initialize s = s by MEMSTR_0:44; set b = DataLoc ((s . a),i); set WHL = while<0 (a,i,I); set pWHL = stop (while<0 (a,i,I)); set P1 = P +* (stop (while<0 (a,i,I))); set i1 = (a,i) >=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); assume A2: s . (DataLoc ((s . a),i)) < 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 & not ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) or IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) set Es = IExec (I,P,s); set bj = DataLoc (((Initialize (IExec (I,P,s))) . a),i); set EP = P; set PI = P +* (stop I); set m1 = (LifeSpan ((P +* (stop I)),s)) + 2; set s2 = Initialize (IExec (I,P,s)); set P2 = P +* (stop (while<0 (a,i,I))); set m2 = LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))); assume A6: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) > t . (DataLoc ((s . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_thesis: IExec ((while<0 (a,i,I)),P,s) = IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) then while<0 (a,i,I) is_halting_on s,P by Th15; then A7: P +* (stop (while<0 (a,i,I))) halts_on s by I, SCMPDS_6:def_3; A8: stop I c= P +* (stop I) by FUNCT_4:25; A9: for x being Int_position st x in X holds s . x = s . x ; then I is_halting_on s,P by A2, A6; then A10: P +* (stop I) halts_on s by I, SCMPDS_6:def_3; (P +* (stop I)) +* (stop I) halts_on s by A10; then A12: I is_halting_on s,P +* (stop I) by I, SCMPDS_6:def_3; A13: (Initialize (IExec (I,P,s))) . a = (IExec (I,P,s)) . a by SCMPDS_5:15 .= s . a by A2, A6, A9 ; now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_(_for_x_being_Int_position_st_x_in_X_holds_ t_._x_=_(Initialize_(IExec_(I,P,s)))_._x_)_&_t_._a_=_(Initialize_(IExec_(I,P,s)))_._a_&_t_._(DataLoc_(((Initialize_(IExec_(I,P,s)))_._a),i))_<_0_holds_ (_(IExec_(I,Q,t))_._a_=_t_._a_&_(IExec_(I,Q,t))_._(DataLoc_(((Initialize_(IExec_(I,P,s)))_._a),i))_>_t_._(DataLoc_(((Initialize_(IExec_(I,P,s)))_._a),i))_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_(_for_x_being_Int_position_st_x_in_X_holds_ (IExec_(I,Q,t))_._x_=_t_._x_)_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) < 0 holds ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( ( for x being Int_position st x in X holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) < 0 implies ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) assume that A14: for x being Int_position st x in X holds t . x = (Initialize (IExec (I,P,s))) . x and A15: ( t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) < 0 ) ; ::_thesis: ( (IExec (I,Q,t)) . a = t . a & (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) A16: now__::_thesis:_for_x_being_Int_position_st_x_in_X_holds_ t_._x_=_s_._x let x be Int_position; ::_thesis: ( x in X implies t . x = s . x ) assume A17: x in X ; ::_thesis: t . x = s . x hence t . x = (Initialize (IExec (I,P,s))) . x by A14 .= (IExec (I,P,s)) . x by SCMPDS_5:15 .= s . x by A2, A6, A9, A17 ; ::_thesis: verum end; hence (IExec (I,Q,t)) . a = t . a by A6, A13, A15; ::_thesis: ( (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) & I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) thus (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) by A6, A13, A15, A16; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) thus ( I is_closed_on t,Q & I is_halting_on t,Q & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) by A6, A13, A15, A16; ::_thesis: verum end; then while<0 (a,i,I) is_halting_on Initialize (IExec (I,P,s)),P by Th15; then XX: P +* (stop (while<0 (a,i,I))) halts_on Initialize (Initialize (IExec (I,P,s))) by SCMPDS_6:def_3; set m0 = LifeSpan ((P +* (stop (while<0 (a,i,I)))),s); set s4 = Comput ((P +* (stop (while<0 (a,i,I)))),s,1); set P4 = P +* (stop (while<0 (a,i,I))); A21: IC s = 0 by MEMSTR_0:def_11; A22: while<0 (a,i,I) = ((a,i) >=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; A23: Comput ((P +* (stop (while<0 (a,i,I)))),s,(0 + 1)) = Following ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,0))) by EXTPRO_1:3 .= Following ((P +* (stop (while<0 (a,i,I)))),s) .= Exec (((a,i) >=0_goto ((card I) + 2)),s) by A22, I, SCMPDS_6:11 ; A26: IC (Comput ((P +* (stop (while<0 (a,i,I)))),s,1)) = succ (IC s) by A2, A23, SCMPDS_2:57 .= 0 + 1 by A21 ; set mI = LifeSpan ((P +* (stop I)),s); set s5 = Comput ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s))); set P5 = P +* (stop (while<0 (a,i,I))); set l1 = (card I) + 1; for a being Int_position holds s . a = (Comput ((P +* (stop (while<0 (a,i,I)))),s,1)) . a by A23, SCMPDS_2:57; then A28: DataPart s = DataPart (Comput ((P +* (stop (while<0 (a,i,I)))),s,1)) by SCMPDS_4:8; set m3 = (LifeSpan ((P +* (stop I)),s)) + 1; set s6 = Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)); set P6 = P +* (stop (while<0 (a,i,I))); (card I) + 1 < (card I) + 2 by XREAL_1:6; then A29: (card I) + 1 in dom (while<0 (a,i,I)) by Th7; set s7 = Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1)); YY: while<0 (a,i,I) c= stop (while<0 (a,i,I)) by AFINSQ_1:74; stop (while<0 (a,i,I)) c= P +* (stop (while<0 (a,i,I))) by FUNCT_4:25; then A31: while<0 (a,i,I) c= P +* (stop (while<0 (a,i,I))) by YY, XBOOLE_1:1; Shift (I,1) c= while<0 (a,i,I) by Lm2; then A32: Shift (I,1) c= P +* (stop (while<0 (a,i,I))) by A31, XBOOLE_1:1; A33: I is_closed_on s,P +* (stop I) by A2, A6, A9; then A34: IC (Comput ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s)))) = (card I) + 1 by A8, A12, A26, A28, A32, SCMPDS_7:18; A35: (P +* (stop (while<0 (a,i,I)))) /. (IC (Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) = (P +* (stop (while<0 (a,i,I)))) . (IC (Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) by PBOOLE:143; A36: Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)) = Comput ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s))) by EXTPRO_1:4; then A37: CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) = (P +* (stop (while<0 (a,i,I)))) . ((card I) + 1) by A8, A12, A33, A26, A28, A32, A35, SCMPDS_7:18 .= (while<0 (a,i,I)) . ((card I) + 1) by A29, A31, GRFUNC_1:2 .= goto (- ((card I) + 1)) by Th8 ; A39: Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1)) = Following ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card I) + 1))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) by A37 ; then IC (Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) = ICplusConst ((Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1))),(0 - ((card I) + 1))) by SCMPDS_2:54 .= 0 by A34, A36, SCMPDS_7:1 ; then A40: IC (Initialize (IExec (I,P,s))) = IC (Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 2))) by MEMSTR_0:def_11; A41: DataPart (Comput ((P +* (stop I)),s,(LifeSpan ((P +* (stop I)),s)))) = DataPart (Comput ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s)))) by A8, A12, A33, A26, A28, A32, SCMPDS_7:18; now__::_thesis:_for_x_being_Int_position_holds_(Comput_((P_+*_(stop_(while<0_(a,i,I)))),s,(((LifeSpan_((P_+*_(stop_I)),s))_+_1)_+_1)))_._x_=_(Initialize_(IExec_(I,P,s)))_._x let x be Int_position; ::_thesis: (Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) . x = (Initialize (IExec (I,P,s))) . x A42: not x in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; (Comput ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s)))) . x = (Comput ((P +* (stop I)),s,(LifeSpan ((P +* (stop I)),s)))) . x by A41, SCMPDS_4:8 .= (Result ((P +* (stop I)),s)) . x by A10, EXTPRO_1:23 .= (IExec (I,P,s)) . x by SCMPDS_4:def_5 ; hence (Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) . x = (IExec (I,P,s)) . x by A36, A39, SCMPDS_2:54 .= (Initialize (IExec (I,P,s))) . x by A42, FUNCT_4:11 ; ::_thesis: verum end; then A44: DataPart (Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) = DataPart (Initialize (IExec (I,P,s))) by SCMPDS_4:8; A46: Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 2)) = Initialize (IExec (I,P,s)) by A44, A40, MEMSTR_0:78; then CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 2)))) = (a,i) >=0_goto ((card I) + 2) by A22, SCMPDS_6:11; then LifeSpan ((P +* (stop (while<0 (a,i,I)))),s) > (LifeSpan ((P +* (stop I)),s)) + 2 by A7, EXTPRO_1:36, SCMPDS_6:18; then consider nn being Nat such that A47: LifeSpan ((P +* (stop (while<0 (a,i,I)))),s) = ((LifeSpan ((P +* (stop I)),s)) + 2) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 2) + (LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))))))) = Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))),(LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))))) by A46, EXTPRO_1:4; then CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 2) + (LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))))))))) = halt SCMPDS by XX, EXTPRO_1:def_15; then ((LifeSpan ((P +* (stop I)),s)) + 2) + (LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))))) >= LifeSpan ((P +* (stop (while<0 (a,i,I)))),s) by A7, EXTPRO_1:def_15; then A49: LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))) >= nn by A47, XREAL_1:6; A50: Comput ((P +* (stop (while<0 (a,i,I)))),s,(LifeSpan ((P +* (stop (while<0 (a,i,I)))),s))) = Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))),nn) by A46, A47, EXTPRO_1:4; then CurInstr ((P +* (stop (while<0 (a,i,I)))),(Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))),nn))) = halt SCMPDS by A7, EXTPRO_1:def_15; then nn >= LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))) by XX, EXTPRO_1:def_15; then nn = LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))) by A49, XXREAL_0:1; then Result ((P +* (stop (while<0 (a,i,I)))),s) = Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))),(LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))))) by A7, A50, EXTPRO_1:23; hence IExec ((while<0 (a,i,I)),P,s) = Comput ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s))),(LifeSpan ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))))) by SCMPDS_4:def_5 .= Result ((P +* (stop (while<0 (a,i,I)))),(Initialize (IExec (I,P,s)))) by XX, EXTPRO_1:23 .= IExec ((while<0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) by SCMPDS_4:def_5 ; ::_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 :: SCMPDS_8:def 3 (((a,i) <=0_goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 1))); coherence (((a,i) <=0_goto ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 1))) is Program of SCMPDS ; end; :: deftheorem defines while>0 SCMPDS_8: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 ((card I) + 2)) ';' I) ';' (goto (- ((card I) + 1))); 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 ((card I) + 2); reconsider PF = (Load ((a,i) <=0_goto ((card I) + 2))) ';' I as shiftable Program of SCMPDS ; A1: PF = ((a,i) <=0_goto ((card I) + 2)) ';' I by SCMPDS_4:def_2; then card PF = (card I) + 1 by SCMPDS_6:6; then (card PF) + (- ((card I) + 1)) = 0 ; hence while>0 (a,i,I) is shiftable by A1, 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 ; ; end; theorem Th17: :: SCMPDS_8:17 for a being Int_position for i being Integer for I being Program of SCMPDS holds card (while>0 (a,i,I)) = (card I) + 2 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) + 2 let i be Integer; ::_thesis: for I being Program of SCMPDS holds card (while>0 (a,i,I)) = (card I) + 2 let I be Program of SCMPDS; ::_thesis: card (while>0 (a,i,I)) = (card I) + 2 set i1 = (a,i) <=0_goto ((card I) + 2); set I4 = ((a,i) <=0_goto ((card I) + 2)) ';' I; thus card (while>0 (a,i,I)) = (card (((a,i) <=0_goto ((card I) + 2)) ';' I)) + 1 by SCMP_GCD:4 .= ((card I) + 1) + 1 by SCMPDS_6:6 .= (card I) + 2 ; ::_thesis: verum end; Lm3: 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) + 3 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) + 3 let i be Integer; ::_thesis: for I being Program of SCMPDS holds card (stop (while>0 (a,i,I))) = (card I) + 3 let I be Program of SCMPDS; ::_thesis: card (stop (while>0 (a,i,I))) = (card I) + 3 thus card (stop (while>0 (a,i,I))) = (card (while>0 (a,i,I))) + 1 by COMPOS_1:55 .= ((card I) + 2) + 1 by Th17 .= (card I) + 3 ; ::_thesis: verum end; theorem Th18: :: SCMPDS_8:18 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) + 2 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) + 2 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) + 2 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) + 2 iff m in dom (while>0 (a,i,I)) ) let I be Program of SCMPDS; ::_thesis: ( m < (card I) + 2 iff m in dom (while>0 (a,i,I)) ) card (while>0 (a,i,I)) = (card I) + 2 by Th17; hence ( m < (card I) + 2 iff m in dom (while>0 (a,i,I)) ) by AFINSQ_1:66; ::_thesis: verum end; theorem Th19: :: SCMPDS_8:19 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 ((card I) + 2) & (while>0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) 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 ((card I) + 2) & (while>0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) let i be Integer; ::_thesis: for I being Program of SCMPDS holds ( (while>0 (a,i,I)) . 0 = (a,i) <=0_goto ((card I) + 2) & (while>0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) let I be Program of SCMPDS; ::_thesis: ( (while>0 (a,i,I)) . 0 = (a,i) <=0_goto ((card I) + 2) & (while>0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) ) set i1 = (a,i) <=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); set I4 = ((a,i) <=0_goto ((card I) + 2)) ';' I; set J5 = I ';' (goto (- ((card I) + 1))); set WHL = while>0 (a,i,I); while>0 (a,i,I) = ((a,i) <=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; hence (while>0 (a,i,I)) . 0 = (a,i) <=0_goto ((card I) + 2) by SCMPDS_6:7; ::_thesis: (while>0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) card (((a,i) <=0_goto ((card I) + 2)) ';' I) = (card I) + 1 by SCMPDS_6:6; hence (while>0 (a,i,I)) . ((card I) + 1) = goto (- ((card I) + 1)) by SCMP_GCD:6; ::_thesis: verum end; theorem Th20: :: SCMPDS_8:20 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 d1 = 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 i1 = (a,i) <=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); set WHL = while>0 (a,i,I); set pWHL = 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 P4 = P +* (stop (while>0 (a,i,I))); A3: IC (Initialize s) = 0 by MEMSTR_0:def_11; A4: 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 A5: (Initialize s) . (DataLoc (((Initialize s) . a),i)) = (Initialize s) . (DataLoc ((s . a),i)) by FUNCT_4:11 .= s . (DataLoc ((s . a),i)) by A4, FUNCT_4:11 ; A6: while>0 (a,i,I) = ((a,i) <=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; 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 ((card I) + 2)),(Initialize s)) by A6, SCMPDS_6:11 ; then A7: IC (Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize s),1)) = ICplusConst ((Initialize s),((card I) + 2)) by A1, A5, SCMPDS_2:56 .= 0 + ((card I) + 2) by A3, SCMPDS_6:12 ; A8: card (while>0 (a,i,I)) = (card I) + 2 by Th17; then A9: (card I) + 2 in dom (stop (while>0 (a,i,I))) by COMPOS_1:64; stop (while>0 (a,i,I)) c= P +* (stop (while>0 (a,i,I))) by FUNCT_4:25; then (P +* (stop (while>0 (a,i,I)))) . ((card I) + 2) = (stop (while>0 (a,i,I))) . ((card I) + 2) by A9, GRFUNC_1:2 .= halt SCMPDS by A8, COMPOS_1:64 ; then A11: CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize s),1))) = halt SCMPDS by A7, PBOOLE:143; 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))) percases ( 0 < k or k = 0 ) ; suppose 0 < k ; ::_thesis: IC (Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize s),b1)) in dom (stop (while>0 (a,i,I))) then 1 + 0 <= 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 A9, A7, A11, EXTPRO_1:5; ::_thesis: verum end; 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 A3, COMPOS_1:36; ::_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 A11, EXTPRO_1:29; hence while>0 (a,i,I) is_halting_on s,P by SCMPDS_6:def_3; ::_thesis: verum end; theorem Th21: :: SCMPDS_8:21 for P being Instruction-Sequence of SCMPDS for s being 0 -started 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,s) = s +* (Start-At (((card I) + 2),SCMPDS)) proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started 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,s) = s +* (Start-At (((card I) + 2),SCMPDS)) let s be 0 -started 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,s) = s +* (Start-At (((card I) + 2),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,s) = s +* (Start-At (((card I) + 2),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,s) = s +* (Start-At (((card I) + 2),SCMPDS)) let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) <= 0 implies IExec ((while>0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) ) set d1 = DataLoc ((s . a),i); set WHL = while>0 (a,i,I); set pWHL = stop (while>0 (a,i,I)); set P3 = P +* (stop (while>0 (a,i,I))); set s4 = Comput ((P +* (stop (while>0 (a,i,I)))),s,1); set P4 = P +* (stop (while>0 (a,i,I))); set i1 = (a,i) <=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); set SAl = Start-At (((card I) + 2),SCMPDS); I: Initialize s = s by MEMSTR_0:44; A2: IC s = 0 by MEMSTR_0:def_11; A3: while>0 (a,i,I) = ((a,i) <=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; A6: Comput ((P +* (stop (while>0 (a,i,I)))),s,(0 + 1)) = Following ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,0))) by EXTPRO_1:3 .= Following ((P +* (stop (while>0 (a,i,I)))),s) .= Exec (((a,i) <=0_goto ((card I) + 2)),s) by A3, I, SCMPDS_6:11 ; A7: stop (while>0 (a,i,I)) c= P +* (stop (while>0 (a,i,I))) by FUNCT_4:25; A8: IExec ((while>0 (a,i,I)),P,s) = Result ((P +* (stop (while>0 (a,i,I)))),s) by SCMPDS_4:def_5; assume s . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: IExec ((while>0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) then A9: IC (Comput ((P +* (stop (while>0 (a,i,I)))),s,1)) = ICplusConst (s,((card I) + 2)) by A6, SCMPDS_2:56 .= 0 + ((card I) + 2) by A2, SCMPDS_6:12 ; A12: card (while>0 (a,i,I)) = (card I) + 2 by Th17; then (card I) + 2 in dom (stop (while>0 (a,i,I))) by COMPOS_1:64; then (P +* (stop (while>0 (a,i,I)))) . ((card I) + 2) = (stop (while>0 (a,i,I))) . ((card I) + 2) by A7, GRFUNC_1:2 .= halt SCMPDS by A12, COMPOS_1:64 ; then A13: CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,1))) = halt SCMPDS by A9, PBOOLE:143; then A14: P +* (stop (while>0 (a,i,I))) halts_on s by EXTPRO_1:29; A15: CurInstr ((P +* (stop (while>0 (a,i,I)))),s) = (a,i) <=0_goto ((card I) + 2) by A3, I, SCMPDS_6:11; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_0_+_1_holds_ CurInstr_((P_+*_(stop_(while>0_(a,i,I)))),(Comput_((P_+*_(stop_(while>0_(a,i,I)))),s,l)))_<>_halt_SCMPDS let l be Element of NAT ; ::_thesis: ( l < 0 + 1 implies CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,l))) <> halt SCMPDS ) A16: Comput ((P +* (stop (while>0 (a,i,I)))),s,0) = s ; assume l < 0 + 1 ; ::_thesis: CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,l))) <> halt SCMPDS then l = 0 by NAT_1:13; hence CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,l))) <> halt SCMPDS by A15, A16; ::_thesis: verum end; then for l being Element of NAT st CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,l))) = halt SCMPDS holds 1 <= l ; then B17: LifeSpan ((P +* (stop (while>0 (a,i,I)))),s) = 1 by A13, A14, EXTPRO_1:def_15; A19: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((while>0_(a,i,I)),P,s))_holds_ (IExec_((while>0_(a,i,I)),P,s))_._x_=_(s_+*_(Start-At_(((card_I)_+_2),SCMPDS)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((while>0 (a,i,I)),P,s)) implies (IExec ((while>0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 ) A20: dom (Start-At (((card I) + 2),SCMPDS)) = {(IC )} by FUNCOP_1:13; assume A21: x in dom (IExec ((while>0 (a,i,I)),P,s)) ; ::_thesis: (IExec ((while>0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 percases ( x is Int_position or x = IC ) by A21, SCMPDS_4:6; supposeA22: x is Int_position ; ::_thesis: (IExec ((while>0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 then x <> IC by SCMPDS_2:43; then A23: not x in dom (Start-At (((card I) + 2),SCMPDS)) by A20, TARSKI:def_1; thus (IExec ((while>0 (a,i,I)),P,s)) . x = (Comput ((P +* (stop (while>0 (a,i,I)))),s,1)) . x by A8, B17, A14, EXTPRO_1:23 .= s . x by A6, A22, SCMPDS_2:56 .= (s +* (Start-At (((card I) + 2),SCMPDS))) . x by A23, FUNCT_4:11 ; ::_thesis: verum end; supposeA24: x = IC ; ::_thesis: (IExec ((while>0 (a,i,I)),P,s)) . b1 = (s +* (Start-At (((card I) + 2),SCMPDS))) . b1 hence (IExec ((while>0 (a,i,I)),P,s)) . x = (card I) + 2 by A8, A9, B17, A14, EXTPRO_1:23 .= (s +* (Start-At (((card I) + 2),SCMPDS))) . x by A24, FUNCT_4:113 ; ::_thesis: verum end; end; end; dom (IExec ((while>0 (a,i,I)),P,s)) = the carrier of SCMPDS by PARTFUN1:def_2 .= dom (s +* (Start-At (((card I) + 2),SCMPDS))) by PARTFUN1:def_2 ; hence IExec ((while>0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) by A19, FUNCT_1:2; ::_thesis: verum end; theorem :: SCMPDS_8:22 for P being Instruction-Sequence of SCMPDS for s being 0 -started 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,s)) = (card I) + 2 proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started 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,s)) = (card I) + 2 let s be 0 -started 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,s)) = (card I) + 2 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,s)) = (card I) + 2 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,s)) = (card I) + 2 let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) <= 0 implies IC (IExec ((while>0 (a,i,I)),P,s)) = (card I) + 2 ) assume s . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: IC (IExec ((while>0 (a,i,I)),P,s)) = (card I) + 2 then IExec ((while>0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) by Th21; hence IC (IExec ((while>0 (a,i,I)),P,s)) = (card I) + 2 by FUNCT_4:113; ::_thesis: verum end; theorem :: SCMPDS_8:23 for P being Instruction-Sequence of SCMPDS for s being 0 -started 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,s)) . b = s . b proof let P be Instruction-Sequence of SCMPDS; ::_thesis: for s being 0 -started 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,s)) . b = s . b let s be 0 -started 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,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,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,s)) . b = s . b let i be Integer; ::_thesis: ( s . (DataLoc ((s . a),i)) <= 0 implies (IExec ((while>0 (a,i,I)),P,s)) . b = s . b ) assume s . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: (IExec ((while>0 (a,i,I)),P,s)) . b = s . b then A1: IExec ((while>0 (a,i,I)),P,s) = s +* (Start-At (((card I) + 2),SCMPDS)) by Th21; not b in dom (Start-At (((card I) + 2),SCMPDS)) by SCMPDS_4:18; hence (IExec ((while>0 (a,i,I)),P,s)) . b = s . b by A1, FUNCT_4:11; ::_thesis: verum end; Lm4: for I being Program of SCMPDS for a being Int_position for i being Integer holds Shift (I,1) 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,1) c= while>0 (a,i,I) let a be Int_position; ::_thesis: for i being Integer holds Shift (I,1) c= while>0 (a,i,I) let i be Integer; ::_thesis: Shift (I,1) c= while>0 (a,i,I) set i1 = (a,i) <=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); A1: while>0 (a,i,I) = (((a,i) <=0_goto ((card I) + 2)) ';' I) ';' (Load (goto (- ((card I) + 1)))) by SCMPDS_4:def_3 .= ((Load ((a,i) <=0_goto ((card I) + 2))) ';' I) ';' (Load (goto (- ((card I) + 1)))) by SCMPDS_4:def_2 ; card (Load ((a,i) <=0_goto ((card I) + 2))) = 1 by COMPOS_1:54; hence Shift (I,1) c= while>0 (a,i,I) by A1, SCMPDS_7:3; ::_thesis: verum end; scheme :: SCMPDS_8:sch 3 WhileGHalt{ 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[ State of SCMPDS] } : ( 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 set i1 = (F5(),F6()) <=0_goto ((card F4()) + 2); set i2 = goto (- ((card F4()) + 1)); set WHL = while>0 (F5(),F6(),F4()); set pWHL = 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(t) <= $1 & P1[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,_Q_being_Instruction-Sequence_of_SCMPDS_st_F1(t)_<=_k_+_1_&_P1[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, Q being Instruction-Sequence of SCMPDS st F1(t) <= k + 1 & P1[t] & t . F5() = F2() . F5() holds ( while>0 (F5(),F6(),F4()) is_closed_on b3,b5 & while>0 (F5(),F6(),F4()) is_halting_on b3,b5 ) let Q, Q be Instruction-Sequence of SCMPDS; ::_thesis: ( F1(t) <= k + 1 & P1[t] & t . F5() = F2() . F5() implies ( while>0 (F5(),F6(),F4()) is_closed_on b1,b3 & while>0 (F5(),F6(),F4()) is_halting_on b1,b3 ) ) assume A7: F1(t) <= k + 1 ; ::_thesis: ( P1[t] & t . F5() = F2() . F5() implies ( while>0 (F5(),F6(),F4()) is_closed_on b1,b3 & while>0 (F5(),F6(),F4()) is_halting_on b1,b3 ) ) assume A8: P1[t] ; ::_thesis: ( t . F5() = F2() . F5() implies ( while>0 (F5(),F6(),F4()) is_closed_on b1,b3 & while>0 (F5(),F6(),F4()) is_halting_on b1,b3 ) ) assume A9: t . F5() = F2() . F5() ; ::_thesis: ( while>0 (F5(),F6(),F4()) is_closed_on b1,b3 & while>0 (F5(),F6(),F4()) is_halting_on b1,b3 ) 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,b3 & while>0 (F5(),F6(),F4()) is_halting_on b1,b3 ) hence ( while>0 (F5(),F6(),F4()) is_closed_on t,Q & while>0 (F5(),F6(),F4()) is_halting_on t,Q ) by A9, Th20; ::_thesis: verum end; supposeA10: t . (DataLoc ((F2() . F5()),F6())) > 0 ; ::_thesis: ( while>0 (F5(),F6(),F4()) is_closed_on b1,b3 & while>0 (F5(),F6(),F4()) is_halting_on b1,b3 ) A14: 0 in dom (stop (while>0 (F5(),F6(),F4()))) by COMPOS_1:36; A17: while>0 (F5(),F6(),F4()) = ((F5(),F6()) <=0_goto ((card F4()) + 2)) ';' (F4() ';' (goto (- ((card F4()) + 1)))) by SCMPDS_4:15; set t2 = t; set Q2 = Q +* (stop F4()); set t3 = t; set Q3 = Q +* (stop (while>0 (F5(),F6(),F4()))); set t4 = Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),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())),t); set t5 = Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),t))); set Q5 = Q +* (stop (while>0 (F5(),F6(),F4()))); set l1 = (card F4()) + 1; I: Initialize t = t by MEMSTR_0:44; A21: IC t = 0 by MEMSTR_0:def_11; set m3 = (LifeSpan ((Q +* (stop F4())),t)) + 1; set t6 = Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)); set Q6 = Q +* (stop (while>0 (F5(),F6(),F4()))); set t7 = Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)); set Q7 = Q +* (stop (while>0 (F5(),F6(),F4()))); (card F4()) + 1 < (card F4()) + 2 by XREAL_1:6; then A22: (card F4()) + 1 in dom (while>0 (F5(),F6(),F4())) by Th18; XX: while>0 (F5(),F6(),F4()) c= stop (while>0 (F5(),F6(),F4())) by AFINSQ_1:74; stop (while>0 (F5(),F6(),F4())) c= Q +* (stop (while>0 (F5(),F6(),F4()))) by FUNCT_4:25; then A23: while>0 (F5(),F6(),F4()) c= Q +* (stop (while>0 (F5(),F6(),F4()))) by XX, XBOOLE_1:1; Shift (F4(),1) c= while>0 (F5(),F6(),F4()) by Lm4; then A24: Shift (F4(),1) c= Q +* (stop (while>0 (F5(),F6(),F4()))) by A23, XBOOLE_1:1; A25: Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(0 + 1)) = Following ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,0))) by EXTPRO_1:3 .= Following ((Q +* (stop (while>0 (F5(),F6(),F4())))),t) .= Exec (((F5(),F6()) <=0_goto ((card F4()) + 2)),t) by A17, I, SCMPDS_6:11 ; for a being Int_position holds t . a = (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)) . a by A25, SCMPDS_2:56; then A27: DataPart t = DataPart (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)) by SCMPDS_4:8; F4() is_halting_on t,Q by A4, A8, A9, A10; then A28: Q +* (stop F4()) halts_on t by I, SCMPDS_6:def_3; (Q +* (stop F4())) +* (stop F4()) halts_on t by A28; then A30: F4() is_halting_on t,Q +* (stop F4()) by I, SCMPDS_6:def_3; A31: IExec (F4(),Q,t) = Result ((Q +* (stop F4())),t) by SCMPDS_4:def_5; A32: P1[ Initialize (IExec (F4(),Q,t))] by A4, A8, A9, A10; A33: F4() is_closed_on t,Q by A4, A8, A9, A10; A34: F4() is_closed_on t,Q +* (stop F4()) by A4, A8, A9, A10; A35: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)) = succ (IC t) by A10, A25, A9, SCMPDS_2:56 .= 0 + 1 by A21 ; then A36: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),t)))) = (card F4()) + 1 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())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) = (Q +* (stop (while>0 (F5(),F6(),F4())))) . (IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) by PBOOLE:143; A38: Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)) = Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),t))) by EXTPRO_1:4; then A39: CurInstr ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) = (Q +* (stop (while>0 (F5(),F6(),F4())))) . ((card F4()) + 1) by A20, A30, A34, A35, A27, A24, A37, SCMPDS_7:18 .= (while>0 (F5(),F6(),F4())) . ((card F4()) + 1) by A22, A23, GRFUNC_1:2 .= goto (- ((card F4()) + 1)) by Th19 ; A41: Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)) = Following ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card F4()) + 1))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1)))) by A39 ; then IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1))),(0 - ((card F4()) + 1))) by SCMPDS_2:54 .= 0 by A36, A38, SCMPDS_7:1 ; then A42: Initialize (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) = Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)) by MEMSTR_0:46; A43: DataPart (Comput ((Q +* (stop F4())),t,(LifeSpan ((Q +* (stop F4())),t)))) = DataPart (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)),(LifeSpan ((Q +* (stop F4())),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())))),t,1)),(LifeSpan ((Q +* (stop F4())),t)))) = DataPart (Result ((Q +* (stop F4())),t)) by A28, EXTPRO_1:23 .= DataPart (IExec (F4(),Q,t)) by SCMPDS_4:def_5 ; InsCode (goto (- ((card F4()) + 1))) = 14 by SCMPDS_2:12; then InsCode (goto (- ((card F4()) + 1))) in {0,4,5,6,14} by ENUMSET1:def_3; then A45: Initialize (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) = Initialize (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1))) by A41, Th3 .= Initialize (IExec (F4(),Q,t)) by A44, A38, MEMSTR_0:80 ; A46: now__::_thesis:_not_F1((Initialize_(Comput_((Q_+*_(stop_(while>0_(F5(),F6(),F4())))),t,(((LifeSpan_((Q_+*_(stop_F4())),t))_+_1)_+_1)))))_>_k F1((Initialize (IExec (F4(),Q,t)))) < F1((Initialize t)) by A4, A8, A9, A10, I; then A47: F1((Initialize (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))))) < k + 1 by A7, A45, I, XXREAL_0:2; assume F1((Initialize (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),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())))),t,1)),(LifeSpan ((Q +* (stop F4())),t)))) . F5() = (Comput ((Q +* (stop F4())),t,(LifeSpan ((Q +* (stop F4())),t)))) . F5() by A43, SCMPDS_4:8 .= (Result ((Q +* (stop F4())),t)) . F5() by A28, EXTPRO_1:23 .= F2() . F5() by A9, A4, A8, A10, A31 ; A50: (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))) . F5() = (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,((LifeSpan ((Q +* (stop F4())),t)) + 1))) . F5() by A41, SCMPDS_2:54 .= F2() . F5() by A48, EXTPRO_1:4 ; then A51: while>0 (F5(),F6(),F4()) is_closed_on Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)),Q +* (stop (while>0 (F5(),F6(),F4()))) by A6, A32, A45, A46, A42; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(while>0_(F5(),F6(),F4())))),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())))),t,b1)) in dom (stop (while>0 (F5(),F6(),F4()))) percases ( k < ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,b1)) in dom (stop (while>0 (F5(),F6(),F4()))) then A52: k <= (LifeSpan ((Q +* (stop F4())),t)) + 1 by INT_1:7; percases ( k <= LifeSpan ((Q +* (stop F4())),t) or k = (LifeSpan ((Q +* (stop F4())),t)) + 1 ) by A52, NAT_1:8; supposeA53: k <= LifeSpan ((Q +* (stop F4())),t) ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,b1)) 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())))),t,k)) in dom (stop (while>0 (F5(),F6(),F4()))) hence IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),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())))),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())),t,kn)) as Element of NAT ; kn < k by A54, XREAL_1:29; then kn < LifeSpan ((Q +* (stop F4())),t) by A53, XXREAL_0:2; then (IC (Comput ((Q +* (stop F4())),t,kn))) + 1 = IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,1)),kn)) by A20, A30, A34, A35, A27, A24, SCMPDS_7:16; then A56: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,k)) = lm + 1 by A54, EXTPRO_1:4; IC (Comput ((Q +* (stop F4())),t,kn)) in dom (stop F4()) by A33, I, 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 + 1 <= (card F4()) + 1 by INT_1:7; (card F4()) + 1 < (card F4()) + 3 by XREAL_1:6; then lm + 1 < (card F4()) + 3 by A57, XXREAL_0:2; then lm + 1 < card (stop (while>0 (F5(),F6(),F4()))) by Lm3; hence IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),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())),t)) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,b1)) in dom (stop (while>0 (F5(),F6(),F4()))) (card F4()) + 1 in dom (stop (while>0 (F5(),F6(),F4()))) by A22, COMPOS_1:62; hence IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),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; suppose k >= ((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,b1)) in dom (stop (while>0 (F5(),F6(),F4()))) then consider nn being Nat such that A59: k = (((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1) + nn by NAT_1:10; A61: nn in NAT by ORDINAL1:def_12; Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,k) = Comput (((Q +* (stop (while>0 (F5(),F6(),F4())))) +* (stop (while>0 (F5(),F6(),F4())))),(Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1))),nn) by A59, A61, EXTPRO_1:4; hence IC (Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,k)) in dom (stop (while>0 (F5(),F6(),F4()))) by A51, A61, A42, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence while>0 (F5(),F6(),F4()) is_closed_on t,Q by I, SCMPDS_6:def_2; ::_thesis: while>0 (F5(),F6(),F4()) is_halting_on t,Q RR: (Q +* (stop (while>0 (F5(),F6(),F4())))) +* (stop (while>0 (F5(),F6(),F4()))) = Q +* (stop (while>0 (F5(),F6(),F4()))) ; while>0 (F5(),F6(),F4()) is_halting_on Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)),Q +* (stop (while>0 (F5(),F6(),F4()))) by A6, A32, A50, A45, A46, A42; then Q +* (stop (while>0 (F5(),F6(),F4()))) halts_on Comput ((Q +* (stop (while>0 (F5(),F6(),F4())))),t,(((LifeSpan ((Q +* (stop F4())),t)) + 1) + 1)) by A42, RR, SCMPDS_6:def_3; then Q +* (stop (while>0 (F5(),F6(),F4()))) halts_on t by EXTPRO_1:22; hence while>0 (F5(),F6(),F4()) is_halting_on t,Q by I, SCMPDS_6:def_3; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; set n = F1(F2()); A62: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st F1(t) <= 0 & P1[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(t) <= 0 & P1[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 A63: F1(t) <= 0 and A64: P1[t] and A65: 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 ) F1(t) = 0 by A63; then t . (DataLoc ((F2() . F5()),F6())) <= 0 by A2, A64; hence ( while>0 (F5(),F6(),F4()) is_closed_on t,Q & while>0 (F5(),F6(),F4()) is_halting_on t,Q ) by A65, Th20; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A62, 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; ::_thesis: verum end; scheme :: SCMPDS_8:sch 4 WhileGExec{ 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[ State of SCMPDS] } : 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 set WHL = while>0 (F5(),F6(),F4()); set pWHL = stop (while>0 (F5(),F6(),F4())); set P1 = F3() +* (stop (while>0 (F5(),F6(),F4()))); set PI = F3() +* (stop F4()); set m1 = (LifeSpan ((F3() +* (stop F4())),F2())) + 2; set s2 = Initialize (IExec (F4(),F3(),F2())); set P2 = F3() +* (stop (while>0 (F5(),F6(),F4()))); set m2 = LifeSpan ((F3() +* (stop (while>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))); I: Initialize F2() = F2() by MEMSTR_0:44; A9: P1[F2()] by A4; A10: stop F4() c= F3() +* (stop F4()) by FUNCT_4:25; A11: F4() is_closed_on F2(),F3() +* (stop F4()) by A2, A4, A5; F4() is_halting_on F2(),F3() by A2, A4, A5; then A12: F3() +* (stop F4()) halts_on F2() by I, SCMPDS_6:def_3; (F3() +* (stop F4())) +* (stop F4()) halts_on F2() by A12; then A14: F4() is_halting_on F2(),F3() +* (stop F4()) by I, SCMPDS_6:def_3; set Es = IExec (F4(),F3(),F2()); set bj = DataLoc (((Initialize (IExec (F4(),F3(),F2()))) . F5()),F6()); set EP = F3(); deffunc H1( State of SCMPDS) -> Element of NAT = F1($1); A15: 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; A16: for t being 0 -started State of SCMPDS st P1[t] & H1(t) = 0 holds t . (DataLoc ((F2() . F5()),F6())) <= 0 by A3; ( while>0 (F5(),F6(),F4()) is_closed_on F2(),F3() & while>0 (F5(),F6(),F4()) is_halting_on F2(),F3() ) from SCMPDS_8:sch_3(A16, A9, A15); then A17: F3() +* (stop (while>0 (F5(),F6(),F4()))) halts_on F2() by I, SCMPDS_6:def_3; deffunc H2( State of SCMPDS) -> Element of NAT = F1($1); A18: (Initialize (IExec (F4(),F3(),F2()))) . F5() = (IExec (F4(),F3(),F2())) . F5() by SCMPDS_5:15 .= F2() . F5() by A2, A4, A5 ; then A19: 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; A20: 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, A18; A21: P1[ Initialize (IExec (F4(),F3(),F2()))] by A2, A4, A5; ( 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 SCMPDS_8:sch_3(A19, A21, A20); then A22: F3() +* (stop (while>0 (F5(),F6(),F4()))) halts_on Initialize (Initialize (IExec (F4(),F3(),F2()))) 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 ((card F4()) + 2); set i2 = goto (- ((card F4()) + 1)); set b = DataLoc ((F2() . F5()),F6()); A23: while>0 (F5(),F6(),F4()) = ((F5(),F6()) <=0_goto ((card F4()) + 2)) ';' (F4() ';' (goto (- ((card F4()) + 1)))) by SCMPDS_4:15; 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 l1 = (card F4()) + 1; A24: IC F2() = 0 by MEMSTR_0:def_11; A25: 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 ((card F4()) + 2)),F2()) by A23, I, SCMPDS_6:11 ; for a being Int_position holds F2() . a = (Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),1)) . a by A25, SCMPDS_2:56; then A27: DataPart F2() = DataPart (Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),1)) by SCMPDS_4:8; 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()) + 1 < (card F4()) + 2 by XREAL_1:6; then A28: (card F4()) + 1 in dom (while>0 (F5(),F6(),F4())) by Th18; set m0 = LifeSpan ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2()); set s7 = Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1)); XX: while>0 (F5(),F6(),F4()) c= stop (while>0 (F5(),F6(),F4())) by AFINSQ_1:74; stop (while>0 (F5(),F6(),F4())) c= F3() +* (stop (while>0 (F5(),F6(),F4()))) by FUNCT_4:25; then A33: while>0 (F5(),F6(),F4()) c= F3() +* (stop (while>0 (F5(),F6(),F4()))) by XX, XBOOLE_1:1; Shift (F4(),1) c= while>0 (F5(),F6(),F4()) by Lm4; then A34: Shift (F4(),1) c= F3() +* (stop (while>0 (F5(),F6(),F4()))) by A33, XBOOLE_1:1; A35: IC (Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),1)) = succ (IC F2()) by A2, A25, SCMPDS_2:56 .= 0 + 1 by A24 ; then A36: 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()) + 1 by A10, A14, A11, A27, A34, SCMPDS_7:18; A37: (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; A38: 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 A39: 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()) + 1) by A10, A14, A11, A35, A27, A34, A37, SCMPDS_7:18 .= (while>0 (F5(),F6(),F4())) . ((card F4()) + 1) by A28, A33, GRFUNC_1:2 .= goto (- ((card F4()) + 1)) by Th19 ; A41: 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()) + 1))),(Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 1)))) by A39 ; 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()) + 1))) by SCMPDS_2:54 .= 0 by A36, A38, SCMPDS_7:1 ; then A42: 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:def_11; A43: 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 A10, A14, A11, A35, A27, A34, 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 A44: 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 A43, SCMPDS_4:8 .= (Result ((F3() +* (stop F4())),F2())) . x by A12, EXTPRO_1:23 .= (IExec (F4(),F3(),F2())) . x by SCMPDS_4:def_5 ; hence (Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),(((LifeSpan ((F3() +* (stop F4())),F2())) + 1) + 1))) . x = (IExec (F4(),F3(),F2())) . x by A38, A41, SCMPDS_2:54 .= (Initialize (IExec (F4(),F3(),F2()))) . x by A44, FUNCT_4:11 ; ::_thesis: verum end; then A46: 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; A48: Comput ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2(),((LifeSpan ((F3() +* (stop F4())),F2())) + 2)) = Initialize (IExec (F4(),F3(),F2())) by A46, A42, 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 ((card F4()) + 2) by A23, SCMPDS_6:11; then LifeSpan ((F3() +* (stop (while>0 (F5(),F6(),F4())))),F2()) > (LifeSpan ((F3() +* (stop F4())),F2())) + 2 by A17, EXTPRO_1:36, SCMPDS_6:17; 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; 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()))))))) = 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()))))))))) = halt SCMPDS by A22, 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 A17, 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 A17, EXTPRO_1:def_15; then nn >= LifeSpan ((F3() +* (stop (while>0 (F5(),F6(),F4())))),(Initialize (IExec (F4(),F3(),F2())))) by A22, 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 A17, 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 A22, EXTPRO_1:23 .= IExec ((while>0 (F5(),F6(),F4())),F3(),(Initialize (IExec (F4(),F3(),F2())))) by SCMPDS_4:def_5 ; ::_thesis: verum end; theorem Th24: :: SCMPDS_8:24 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( 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 being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( 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 being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) let i, c be Integer; ::_thesis: for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) let X, Y be set ; ::_thesis: for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) let f be Function of (product (the_Values_of SCMPDS)),NAT; ::_thesis: ( ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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); set WHL = while>0 (a,i,I); set pWHL = stop (while>0 (a,i,I)); set pI = stop I; set i1 = (a,i) <=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); defpred S1[ Element of NAT ] means for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st f . t <= $1 & ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a holds ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ); assume A2: for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: ( ex x being Int_position st ( x in X & not s . x >= c + (s . (DataLoc ((s . a),i))) ) or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) or ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) ) assume A3: for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) or ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) ) assume A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_thesis: ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) 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_f_._t_<=_k_+_1_&_(_for_x_being_Int_position_st_x_in_X_holds_ t_._x_>=_c_+_(t_._(DataLoc_((s_._a),i)))_)_&_(_for_x_being_Int_position_st_x_in_Y_holds_ t_._x_=_s_._x_)_&_t_._a_=_s_._a_holds_ (_while>0_(a,i,I)_is_closed_on_t,Q_&_while>0_(a,i,I)_is_halting_on_t,Q_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st f . t <= k + 1 & ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a holds ( while>0 (a,i,I) is_closed_on b2,b3 & while>0 (a,i,I) is_halting_on b2,b3 ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( f . t <= k + 1 & ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a implies ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) ) T: Initialize t = t by MEMSTR_0:44; assume A7: f . t <= k + 1 ; ::_thesis: ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a implies ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) ) assume A8: for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ; ::_thesis: ( ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a implies ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) ) assume A9: for x being Int_position st x in Y holds t . x = s . x ; ::_thesis: ( t . a = s . a implies ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) ) assume A10: t . a = s . a ; ::_thesis: ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) percases ( t . (DataLoc ((s . a),i)) <= 0 or t . (DataLoc ((s . a),i)) > 0 ) ; suppose t . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) hence ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) by A10, Th20; ::_thesis: verum end; supposeA11: t . (DataLoc ((s . a),i)) > 0 ; ::_thesis: ( while>0 (a,i,I) is_closed_on b1,b2 & while>0 (a,i,I) is_halting_on b1,b2 ) A15: 0 in dom (stop (while>0 (a,i,I))) by COMPOS_1:36; A18: while>0 (a,i,I) = ((a,i) <=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; set Q2 = Q +* (stop I); set Q3 = Q +* (stop (while>0 (a,i,I))); set t4 = Comput ((Q +* (stop (while>0 (a,i,I)))),t,1); set Q4 = Q +* (stop (while>0 (a,i,I))); A21: stop I c= Q +* (stop I) by FUNCT_4:25; A22: Comput ((Q +* (stop (while>0 (a,i,I)))),t,(0 + 1)) = Following ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,0))) by EXTPRO_1:3 .= Following ((Q +* (stop (while>0 (a,i,I)))),t) .= Exec (((a,i) <=0_goto ((card I) + 2)),t) by A18, T, SCMPDS_6:11 ; for a being Int_position holds t . a = (Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)) . a by A22, SCMPDS_2:56; then A24: DataPart t = DataPart (Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)) by SCMPDS_4:8; XX: while>0 (a,i,I) c= stop (while>0 (a,i,I)) by AFINSQ_1:74; stop (while>0 (a,i,I)) c= Q +* (stop (while>0 (a,i,I))) by FUNCT_4:25; then A25: while>0 (a,i,I) c= Q +* (stop (while>0 (a,i,I))) by XX, XBOOLE_1:1; Shift (I,1) c= while>0 (a,i,I) by Lm4; then A26: Shift (I,1) c= Q +* (stop (while>0 (a,i,I))) by A25, XBOOLE_1:1; A27: IExec (I,Q,t) = Result ((Q +* (stop I)),t) by SCMPDS_4:def_5; set m2 = LifeSpan ((Q +* (stop I)),t); set t5 = Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t))); set Q5 = Q +* (stop (while>0 (a,i,I))); set l1 = (card I) + 1; A28: IC t = 0 by MEMSTR_0:def_11; set m3 = (LifeSpan ((Q +* (stop I)),t)) + 1; set t6 = Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)); set Q6 = Q +* (stop (while>0 (a,i,I))); set t7 = Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)); set Q7 = Q +* (stop (while>0 (a,i,I))); (card I) + 1 < (card I) + 2 by XREAL_1:6; then A29: (card I) + 1 in dom (while>0 (a,i,I)) by Th18; A30: I is_closed_on t,Q by A4, A8, A9, A10, A11; A31: I is_closed_on t,Q +* (stop I) by A4, A8, A9, A10, A11; I is_halting_on t,Q by A4, A8, A9, A10, A11; then A32: Q +* (stop I) halts_on t by T, SCMPDS_6:def_3; (Q +* (stop I)) +* (stop I) halts_on t by A32; then A34: I is_halting_on t,Q +* (stop I) by T, SCMPDS_6:def_3; A35: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)) = succ (IC t) by A11, A22, A10, SCMPDS_2:56 .= 0 + 1 by A28 ; then A36: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) = (card I) + 1 by A21, A34, A31, A24, A26, SCMPDS_7:18; A37: (Q +* (stop (while>0 (a,i,I)))) /. (IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) = (Q +* (stop (while>0 (a,i,I)))) . (IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) by PBOOLE:143; A38: Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)) = Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t))) by EXTPRO_1:4; then A39: CurInstr ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) = (Q +* (stop (while>0 (a,i,I)))) . ((card I) + 1) by A21, A34, A31, A35, A24, A26, A37, SCMPDS_7:18 .= (while>0 (a,i,I)) . ((card I) + 1) by A29, A25, GRFUNC_1:2 .= goto (- ((card I) + 1)) by Th19 ; A41: Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)) = Following ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card I) + 1))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1)))) by A39 ; then IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) = ICplusConst ((Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1))),(0 - ((card I) + 1))) by SCMPDS_2:54 .= 0 by A36, A38, SCMPDS_7:1 ; then A42: Initialize (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) = Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)) by MEMSTR_0:46; A43: DataPart (Comput ((Q +* (stop I)),t,(LifeSpan ((Q +* (stop I)),t)))) = DataPart (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) by A21, A34, A31, A35, A24, A26, SCMPDS_7:18; then A44: DataPart (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) = DataPart (Result ((Q +* (stop I)),t)) by A32, EXTPRO_1:23 .= DataPart (IExec (I,Q,t)) by SCMPDS_4:def_5 ; A45: now__::_thesis:_for_x_being_Int_position_st_x_in_Y_holds_ (Comput_((Q_+*_(stop_(while>0_(a,i,I)))),t,(((LifeSpan_((Q_+*_(stop_I)),t))_+_1)_+_1)))_._x_=_s_._x let x be Int_position; ::_thesis: ( x in Y implies (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = s . x ) assume A46: x in Y ; ::_thesis: (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = s . x thus (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) . x by A38, A41, SCMPDS_2:54 .= (IExec (I,Q,t)) . x by A44, SCMPDS_3:3 .= t . x by A4, A8, A9, A10, A11, A46 .= s . x by A9, A46 ; ::_thesis: verum end; InsCode (goto (- ((card I) + 1))) = 14 by SCMPDS_2:12; then InsCode (goto (- ((card I) + 1))) in {0,4,5,6,14} by ENUMSET1:def_3; then A47: Initialize (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) = Initialize (Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1))) by A41, Th3 .= Initialize (IExec (I,Q,t)) by A44, A38, MEMSTR_0:80 ; A48: now__::_thesis:_not_f_._(Initialize_(Comput_((Q_+*_(stop_(while>0_(a,i,I)))),t,(((LifeSpan_((Q_+*_(stop_I)),t))_+_1)_+_1))))_>_k f . (Initialize (IExec (I,Q,t))) < f . (Initialize t) by A4, A8, A9, A10, A11, T; then A49: f . (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) < k + 1 by A7, A47, A42, T, XXREAL_0:2; assume f . (Initialize (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)))) > k ; ::_thesis: contradiction hence contradiction by A49, A42, INT_1:7; ::_thesis: verum end; A50: (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . (DataLoc ((s . a),i)) = (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) . (DataLoc ((s . a),i)) by A38, A41, SCMPDS_2:54 .= (IExec (I,Q,t)) . (DataLoc ((s . a),i)) by A44, SCMPDS_3:3 ; A51: now__::_thesis:_for_x_being_Int_position_st_x_in_X_holds_ (Comput_((Q_+*_(stop_(while>0_(a,i,I)))),t,(((LifeSpan_((Q_+*_(stop_I)),t))_+_1)_+_1)))_._x_>=_c_+_((Comput_((Q_+*_(stop_(while>0_(a,i,I)))),t,(((LifeSpan_((Q_+*_(stop_I)),t))_+_1)_+_1)))_._(DataLoc_((s_._a),i))) let x be Int_position; ::_thesis: ( x in X implies (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x >= c + ((Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . (DataLoc ((s . a),i))) ) assume A52: x in X ; ::_thesis: (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x >= c + ((Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . (DataLoc ((s . a),i))) (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x = (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) . x by A38, A41, SCMPDS_2:54 .= (IExec (I,Q,t)) . x by A44, SCMPDS_3:3 ; hence (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . x >= c + ((Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . (DataLoc ((s . a),i))) by A4, A8, A9, A10, A11, A50, A52; ::_thesis: verum end; A53: (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),(LifeSpan ((Q +* (stop I)),t)))) . a = (Comput ((Q +* (stop I)),t,(LifeSpan ((Q +* (stop I)),t)))) . a by A43, SCMPDS_4:8 .= (Result ((Q +* (stop I)),t)) . a by A32, EXTPRO_1:23 .= s . a by A10, A4, A8, A9, A11, A27 ; A55: (Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))) . a = (Comput ((Q +* (stop (while>0 (a,i,I)))),t,((LifeSpan ((Q +* (stop I)),t)) + 1))) . a by A41, SCMPDS_2:54 .= s . a by A53, EXTPRO_1:4 ; then A56: while>0 (a,i,I) is_closed_on Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)),Q +* (stop (while>0 (a,i,I))) by A6, A51, A45, A48, A42; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((Q_+*_(stop_(while>0_(a,i,I)))),t,k))_in_dom_(stop_(while>0_(a,i,I))) let k be Element of NAT ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,b1)) in dom (stop (while>0 (a,i,I))) percases ( k < ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 or k >= ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 ) ; suppose k < ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,b1)) in dom (stop (while>0 (a,i,I))) then A57: k <= (LifeSpan ((Q +* (stop I)),t)) + 1 by INT_1:7; hereby ::_thesis: verum percases ( k <= LifeSpan ((Q +* (stop I)),t) or k = (LifeSpan ((Q +* (stop I)),t)) + 1 ) by A57, NAT_1:8; supposeA58: k <= LifeSpan ((Q +* (stop I)),t) ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) hereby ::_thesis: verum percases ( k = 0 or k <> 0 ) ; suppose k = 0 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) hence IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) by A15, A28, EXTPRO_1:2; ::_thesis: verum end; suppose k <> 0 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) then consider kn being Nat such that A59: k = kn + 1 by NAT_1:6; reconsider kn = kn as Element of NAT by ORDINAL1:def_12; reconsider lm = IC (Comput ((Q +* (stop I)),t,kn)) as Element of NAT ; kn < k by A59, XREAL_1:29; then kn < LifeSpan ((Q +* (stop I)),t) by A58, XXREAL_0:2; then (IC (Comput ((Q +* (stop I)),t,kn))) + 1 = IC (Comput ((Q +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,1)),kn)) by A21, A34, A31, A35, A24, A26, SCMPDS_7:16; then A61: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) = lm + 1 by A59, EXTPRO_1:4; IC (Comput ((Q +* (stop I)),t,kn)) in dom (stop I) by A30, T, SCMPDS_6:def_2; then lm < card (stop I) by AFINSQ_1:66; then lm < (card I) + 1 by COMPOS_1:55; then A62: lm + 1 <= (card I) + 1 by INT_1:7; (card I) + 1 < (card I) + 3 by XREAL_1:6; then lm + 1 < (card I) + 3 by A62, XXREAL_0:2; then lm + 1 < card (stop (while>0 (a,i,I))) by Lm3; hence IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) by A61, AFINSQ_1:66; ::_thesis: verum end; end; end; end; supposeA63: k = (LifeSpan ((Q +* (stop I)),t)) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) (card I) + 1 in dom (stop (while>0 (a,i,I))) by A29, COMPOS_1:62; hence IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) by A21, A34, A31, A35, A24, A26, A38, A63, SCMPDS_7:18; ::_thesis: verum end; end; end; end; suppose k >= ((LifeSpan ((Q +* (stop I)),t)) + 1) + 1 ; ::_thesis: IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,b1)) in dom (stop (while>0 (a,i,I))) then consider nn being Nat such that A64: k = (((LifeSpan ((Q +* (stop I)),t)) + 1) + 1) + nn by NAT_1:10; A66: nn in NAT by ORDINAL1:def_12; Comput ((Q +* (stop (while>0 (a,i,I)))),t,k) = Comput (((Q +* (stop (while>0 (a,i,I)))) +* (stop (while>0 (a,i,I)))),(Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1))),nn) by A64, A66, EXTPRO_1:4; hence IC (Comput ((Q +* (stop (while>0 (a,i,I)))),t,k)) in dom (stop (while>0 (a,i,I))) by A56, A66, A42, SCMPDS_6:def_2; ::_thesis: verum end; end; end; hence while>0 (a,i,I) is_closed_on t,Q by T, SCMPDS_6:def_2; ::_thesis: while>0 (a,i,I) is_halting_on t,Q RR: (Q +* (stop (while>0 (a,i,I)))) +* (stop (while>0 (a,i,I))) = Q +* (stop (while>0 (a,i,I))) ; while>0 (a,i,I) is_halting_on Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)),Q +* (stop (while>0 (a,i,I))) by A6, A55, A51, A45, A48, A42; then Q +* (stop (while>0 (a,i,I))) halts_on Comput ((Q +* (stop (while>0 (a,i,I)))),t,(((LifeSpan ((Q +* (stop I)),t)) + 1) + 1)) by A42, RR, SCMPDS_6:def_3; then Q +* (stop (while>0 (a,i,I))) halts_on t by EXTPRO_1:22; hence while>0 (a,i,I) is_halting_on t,Q by T, SCMPDS_6:def_3; ::_thesis: verum end; end; end; hence S1[k + 1] ; ::_thesis: verum end; set n = f . s; A67: for x being Int_position st x in Y holds s . x = s . x ; A68: S1[ 0 ] proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st f . t <= 0 & ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a holds ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( f . t <= 0 & ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a implies ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) ) assume f . t <= 0 ; ::_thesis: ( ex x being Int_position st ( x in X & not t . x >= c + (t . (DataLoc ((s . a),i))) ) or ex x being Int_position st ( x in Y & not t . x = s . x ) or not t . a = s . a or ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) ) then f . t = 0 ; then A69: t . (DataLoc ((s . a),i)) <= 0 by A2; assume for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ; ::_thesis: ( ex x being Int_position st ( x in Y & not t . x = s . x ) or not t . a = s . a or ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) ) assume for x being Int_position st x in Y holds t . x = s . x ; ::_thesis: ( not t . a = s . a or ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) ) assume t . a = s . a ; ::_thesis: ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) hence ( while>0 (a,i,I) is_closed_on t,Q & while>0 (a,i,I) is_halting_on t,Q ) by A69, Th20; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A68, A5); then S1[f . s] ; hence ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) by A3, A67; ::_thesis: verum end; theorem Th25: :: SCMPDS_8:25 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds 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 being Int_position for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let a be Int_position; ::_thesis: for i, c being Integer for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let i, c be Integer; ::_thesis: for X, Y being set for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let X, Y be set ; ::_thesis: for f being Function of (product (the_Values_of SCMPDS)),NAT st s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) holds IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) let f be Function of (product (the_Values_of SCMPDS)),NAT; ::_thesis: ( s . (DataLoc ((s . a),i)) > 0 & ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) implies IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) I: Initialize s = s by MEMSTR_0:44; set b = DataLoc ((s . a),i); set WHL = while>0 (a,i,I); set pWHL = stop (while>0 (a,i,I)); set P1 = P +* (stop (while>0 (a,i,I))); set i1 = (a,i) <=0_goto ((card I) + 2); set i2 = goto (- ((card I) + 1)); assume A2: s . (DataLoc ((s . a),i)) > 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS st ( f . t = 0 & not t . (DataLoc ((s . a),i)) <= 0 ) or ex x being Int_position st ( x in X & not s . x >= c + (s . (DataLoc ((s . a),i))) ) or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) or IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) set s4 = Comput ((P +* (stop (while>0 (a,i,I)))),s,1); set P4 = P +* (stop (while>0 (a,i,I))); A3: IC s = 0 by MEMSTR_0:def_11; A4: while>0 (a,i,I) = ((a,i) <=0_goto ((card I) + 2)) ';' (I ';' (goto (- ((card I) + 1)))) by SCMPDS_4:15; A5: Comput ((P +* (stop (while>0 (a,i,I)))),s,(0 + 1)) = Following ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,0))) by EXTPRO_1:3 .= Following ((P +* (stop (while>0 (a,i,I)))),s) .= Exec (((a,i) <=0_goto ((card I) + 2)),s) by A4, I, SCMPDS_6:11 ; set m0 = LifeSpan ((P +* (stop (while>0 (a,i,I)))),s); set Es = IExec (I,P,s); set bj = DataLoc (((Initialize (IExec (I,P,s))) . a),i); set EP = P; assume A9: for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: ( ex x being Int_position st ( x in X & not s . x >= c + (s . (DataLoc ((s . a),i))) ) or ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) or IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) assume A10: for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) or IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ) assume A11: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_thesis: IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) then while>0 (a,i,I) is_halting_on s,P by A9, A10, Th24; then A12: P +* (stop (while>0 (a,i,I))) halts_on s by I, SCMPDS_6:def_3; A13: for x being Int_position st x in Y holds s . x = s . x ; A14: DataLoc (((Initialize (IExec (I,P,s))) . a),i) = DataLoc (((IExec (I,P,s)) . a),i) by SCMPDS_5:15 .= DataLoc ((s . a),i) by A2, A10, A11, A13 ; set PI = P +* (stop I); set m1 = (LifeSpan ((P +* (stop I)),s)) + 2; set s2 = Initialize (IExec (I,P,s)); set P2 = P +* (stop (while>0 (a,i,I))); set m2 = LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))); A17: stop I c= P +* (stop I) by FUNCT_4:25; A18: (Initialize (IExec (I,P,s))) . a = (IExec (I,P,s)) . a by SCMPDS_5:15 .= s . a by A2, A10, A11, A13 ; A19: now__::_thesis:_for_t_being_0_-started_State_of_SCMPDS for_Q_being_Instruction-Sequence_of_SCMPDS_st_(_for_x_being_Int_position_st_x_in_X_holds_ t_._x_>=_c_+_(t_._(DataLoc_(((Initialize_(IExec_(I,P,s)))_._a),i)))_)_&_(_for_x_being_Int_position_st_x_in_Y_holds_ t_._x_=_(Initialize_(IExec_(I,P,s)))_._x_)_&_t_._a_=_(Initialize_(IExec_(I,P,s)))_._a_&_t_._(DataLoc_(((Initialize_(IExec_(I,P,s)))_._a),i))_>_0_holds_ (_(Initialize_(IExec_(I,Q,t)))_._a_=_t_._a_&_I_is_closed_on_t,Q_&_I_is_halting_on_t,Q_&_f_._(Initialize_(IExec_(I,Q,t)))_<_f_._t_&_(_for_x_being_Int_position_st_x_in_X_holds_ (Initialize_(IExec_(I,Q,t)))_._x_>=_c_+_((Initialize_(IExec_(I,Q,t)))_._(DataLoc_(((Initialize_(IExec_(I,P,s)))_._a),i)))_)_&_(_for_x_being_Int_position_st_x_in_Y_holds_ (IExec_(I,Q,t))_._x_=_t_._x_)_) let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 holds ( (Initialize (IExec (I,Q,t))) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 implies ( (Initialize (IExec (I,Q,t))) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) assume that A20: for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) and A21: for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x and A22: t . a = (Initialize (IExec (I,P,s))) . a and A23: t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 ; ::_thesis: ( (Initialize (IExec (I,Q,t))) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) A24: now__::_thesis:_for_x_being_Int_position_st_x_in_Y_holds_ t_._x_=_s_._x let x be Int_position; ::_thesis: ( x in Y implies t . x = s . x ) assume A25: x in Y ; ::_thesis: t . x = s . x hence t . x = (Initialize (IExec (I,P,s))) . x by A21 .= (IExec (I,P,s)) . x by SCMPDS_5:15 .= s . x by A2, A10, A11, A13, A25 ; ::_thesis: verum end; thus (Initialize (IExec (I,Q,t))) . a = (IExec (I,Q,t)) . a by SCMPDS_5:15 .= t . a by A11, A18, A20, A22, A23, A24 ; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) A26: t . a = (IExec (I,P,s)) . a by A22, SCMPDS_5:15 .= s . a by A2, A10, A11, A13 ; hence ( I is_closed_on t,Q & I is_halting_on t,Q ) by A11, A20, A22, A23, A24; ::_thesis: ( f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) thus f . (Initialize (IExec (I,Q,t))) < f . t by A11, A20, A22, A23, A24, A26; ::_thesis: ( ( for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) thus for x being Int_position st x in X holds (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ::_thesis: for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x proof let x be Int_position; ::_thesis: ( x in X implies (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) ( (Initialize (IExec (I,Q,t))) . x = (IExec (I,Q,t)) . x & (Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) = (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) ) by SCMPDS_5:15; hence ( x in X implies (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) by A11, A18, A20, A22, A23, A24; ::_thesis: verum end; thus for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x by A11, A20, A22, A23, A24, A26; ::_thesis: verum end; XX: for x being Int_position st x in X holds (Initialize (IExec (I,P,s))) . x >= c + ((Initialize (IExec (I,P,s))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) proof let x be Int_position; ::_thesis: ( x in X implies (Initialize (IExec (I,P,s))) . x >= c + ((Initialize (IExec (I,P,s))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) F: (Initialize (IExec (I,P,s))) . x = (IExec (I,P,s)) . x by SCMPDS_5:15; G: (Initialize (IExec (I,P,s))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) = (IExec (I,P,s)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) by SCMPDS_5:15; assume x in X ; ::_thesis: (Initialize (IExec (I,P,s))) . x >= c + ((Initialize (IExec (I,P,s))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) hence (Initialize (IExec (I,P,s))) . x >= c + ((Initialize (IExec (I,P,s))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) by A2, A10, A11, A13, A18, F, G; ::_thesis: verum end; for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 holds ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) proof let t be 0 -started State of SCMPDS; ::_thesis: for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 holds ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) let Q be Instruction-Sequence of SCMPDS; ::_thesis: ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x ) & t . a = (Initialize (IExec (I,P,s))) . a & t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 implies ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) assume that Z1: for x being Int_position st x in X holds t . x >= c + (t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) and Z2: for x being Int_position st x in Y holds t . x = (Initialize (IExec (I,P,s))) . x and Z3: t . a = (Initialize (IExec (I,P,s))) . a and Z4: t . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) > 0 ; ::_thesis: ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) thus (IExec (I,Q,t)) . a = (Initialize (IExec (I,Q,t))) . a by SCMPDS_5:15 .= t . a by A19, Z1, Z2, Z3, Z4 ; ::_thesis: ( I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) thus ( I is_closed_on t,Q & I is_halting_on t,Q ) by A19, Z1, Z2, Z3, Z4; ::_thesis: ( f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) thus f . (Initialize (IExec (I,Q,t))) < f . t by A19, Z1, Z2, Z3, Z4; ::_thesis: ( ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) thus for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ::_thesis: for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x proof let x be Int_position; ::_thesis: ( x in X implies (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) ( (IExec (I,Q,t)) . x = (Initialize (IExec (I,Q,t))) . x & (IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) = (Initialize (IExec (I,Q,t))) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i)) ) by SCMPDS_5:15; hence ( x in X implies (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc (((Initialize (IExec (I,P,s))) . a),i))) ) by A19, Z1, Z2, Z3, Z4; ::_thesis: verum end; thus for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x by A19, Z1, Z2, Z3, Z4; ::_thesis: verum end; then while>0 (a,i,I) is_halting_on Initialize (IExec (I,P,s)),P by A9, A14, Th24, XX; then B27: P +* (stop (while>0 (a,i,I))) halts_on Initialize (Initialize (IExec (I,P,s))) by SCMPDS_6:def_3; for a being Int_position holds s . a = (Comput ((P +* (stop (while>0 (a,i,I)))),s,1)) . a by A5, SCMPDS_2:56; then A29: DataPart s = DataPart (Comput ((P +* (stop (while>0 (a,i,I)))),s,1)) by SCMPDS_4:8; set mI = LifeSpan ((P +* (stop I)),s); set s5 = Comput ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s))); set P5 = P +* (stop (while>0 (a,i,I))); set l1 = (card I) + 1; XX: while>0 (a,i,I) c= stop (while>0 (a,i,I)) by AFINSQ_1:74; stop (while>0 (a,i,I)) c= P +* (stop (while>0 (a,i,I))) by FUNCT_4:25; then A30: while>0 (a,i,I) c= P +* (stop (while>0 (a,i,I))) by XX, XBOOLE_1:1; set m3 = (LifeSpan ((P +* (stop I)),s)) + 1; set s7 = Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1)); set P7 = P +* (stop (while>0 (a,i,I))); set s6 = Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)); set P6 = P +* (stop (while>0 (a,i,I))); (card I) + 1 < (card I) + 2 by XREAL_1:6; then A32: (card I) + 1 in dom (while>0 (a,i,I)) by Th18; Shift (I,1) c= while>0 (a,i,I) by Lm4; then A33: Shift (I,1) c= P +* (stop (while>0 (a,i,I))) by A30, XBOOLE_1:1; I is_halting_on s,P by A2, A10, A11, A13; then A34: P +* (stop I) halts_on s by I, SCMPDS_6:def_3; (P +* (stop I)) +* (stop I) halts_on s by A34; then A36: I is_halting_on s,P +* (stop I) by I, SCMPDS_6:def_3; A39: IC (Comput ((P +* (stop (while>0 (a,i,I)))),s,1)) = succ (IC s) by A2, A5, SCMPDS_2:56 .= 0 + 1 by A3 ; A40: I is_closed_on s,P +* (stop I) by A2, A10, A11, A13; then A41: IC (Comput ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s)))) = (card I) + 1 by A17, A36, A39, A29, A33, SCMPDS_7:18; A42: (P +* (stop (while>0 (a,i,I)))) /. (IC (Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) = (P +* (stop (while>0 (a,i,I)))) . (IC (Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) by PBOOLE:143; A43: Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)) = Comput ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s))) by EXTPRO_1:4; then A44: CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) = (P +* (stop (while>0 (a,i,I)))) . ((card I) + 1) by A17, A36, A40, A39, A29, A33, A42, SCMPDS_7:18 .= (while>0 (a,i,I)) . ((card I) + 1) by A32, A30, GRFUNC_1:2 .= goto (- ((card I) + 1)) by Th19 ; A46: Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1)) = Following ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) by EXTPRO_1:3 .= Exec ((goto (- ((card I) + 1))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1)))) by A44 ; then IC (Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) = ICplusConst ((Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 1))),(0 - ((card I) + 1))) by SCMPDS_2:54 .= 0 by A41, A43, SCMPDS_7:1 ; then A47: IC (Initialize (IExec (I,P,s))) = IC (Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 2))) by MEMSTR_0:def_11; A48: DataPart (Comput ((P +* (stop I)),s,(LifeSpan ((P +* (stop I)),s)))) = DataPart (Comput ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s)))) by A17, A36, A40, A39, A29, A33, SCMPDS_7:18; now__::_thesis:_for_x_being_Int_position_holds_(Comput_((P_+*_(stop_(while>0_(a,i,I)))),s,(((LifeSpan_((P_+*_(stop_I)),s))_+_1)_+_1)))_._x_=_(Initialize_(IExec_(I,P,s)))_._x let x be Int_position; ::_thesis: (Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) . x = (Initialize (IExec (I,P,s))) . x A49: not x in dom (Start-At (0,SCMPDS)) by SCMPDS_4:18; (Comput ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,1)),(LifeSpan ((P +* (stop I)),s)))) . x = (Comput ((P +* (stop I)),s,(LifeSpan ((P +* (stop I)),s)))) . x by A48, SCMPDS_4:8 .= (Result ((P +* (stop I)),s)) . x by A34, EXTPRO_1:23 .= (IExec (I,P,s)) . x by SCMPDS_4:def_5 ; hence (Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) . x = (IExec (I,P,s)) . x by A43, A46, SCMPDS_2:54 .= (Initialize (IExec (I,P,s))) . x by A49, FUNCT_4:11 ; ::_thesis: verum end; then A51: DataPart (Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 1) + 1))) = DataPart (Initialize (IExec (I,P,s))) by SCMPDS_4:8; A53: Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 2)) = Initialize (IExec (I,P,s)) by A51, A47, MEMSTR_0:78; then CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,((LifeSpan ((P +* (stop I)),s)) + 2)))) = (a,i) <=0_goto ((card I) + 2) by A4, SCMPDS_6:11; then LifeSpan ((P +* (stop (while>0 (a,i,I)))),s) > (LifeSpan ((P +* (stop I)),s)) + 2 by A12, EXTPRO_1:36, SCMPDS_6:17; then consider nn being Nat such that A54: LifeSpan ((P +* (stop (while>0 (a,i,I)))),s) = ((LifeSpan ((P +* (stop I)),s)) + 2) + nn by NAT_1:10; reconsider nn = nn as Element of NAT by ORDINAL1:def_12; Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 2) + (LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))))))) = Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))),(LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))))) by A53, EXTPRO_1:4; then CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),s,(((LifeSpan ((P +* (stop I)),s)) + 2) + (LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))))))))) = halt SCMPDS by B27, EXTPRO_1:def_15; then ((LifeSpan ((P +* (stop I)),s)) + 2) + (LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))))) >= LifeSpan ((P +* (stop (while>0 (a,i,I)))),s) by A12, EXTPRO_1:def_15; then A56: LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))) >= nn by A54, XREAL_1:6; A57: Comput ((P +* (stop (while>0 (a,i,I)))),s,(LifeSpan ((P +* (stop (while>0 (a,i,I)))),s))) = Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))),nn) by A53, A54, EXTPRO_1:4; then CurInstr ((P +* (stop (while>0 (a,i,I)))),(Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))),nn))) = halt SCMPDS by A12, EXTPRO_1:def_15; then nn >= LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))) by B27, EXTPRO_1:def_15; then nn = LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))) by A56, XXREAL_0:1; then Result ((P +* (stop (while>0 (a,i,I)))),s) = Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))),(LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))))) by A12, A57, EXTPRO_1:23; hence IExec ((while>0 (a,i,I)),P,s) = Comput ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s))),(LifeSpan ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))))) by SCMPDS_4:def_5 .= Result ((P +* (stop (while>0 (a,i,I)))),(Initialize (IExec (I,P,s)))) by B27, EXTPRO_1:23 .= IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) by SCMPDS_4:def_5 ; ::_thesis: verum end; theorem :: SCMPDS_8:26 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 be Int_position; ::_thesis: for i being Integer for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 Integer; ::_thesis: for X being set for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 X be set ; ::_thesis: for f being Function of (product (the_Values_of SCMPDS)),NAT st ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 f be Function of (product (the_Values_of SCMPDS)),NAT; ::_thesis: ( ( for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 b = DataLoc ((s . a),i); assume A2: for t being 0 -started State of SCMPDS st f . t = 0 holds t . (DataLoc ((s . a),i)) <= 0 ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 A3: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_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)))) ) ) then A4: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {} holds t . x >= 0 + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in {} holds (IExec (I,Q,t)) . x >= 0 + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; A5: for x being Int_position st x in {} holds s . x >= 0 + (s . (DataLoc ((s . a),i))) ; for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {} holds t . x >= 0 + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds t . x = s . x ) & 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 & f . (Initialize (IExec (I,Q,t))) < f . t & ( for x being Int_position st x in {} holds (IExec (I,Q,t)) . x >= 0 + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) by A3; hence ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) by A2, A5, Th24; ::_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 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)))) hence IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) by A2, A5, A4, Th25; ::_thesis: verum end; theorem Th27: :: SCMPDS_8:27 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i, c being Integer for X, Y being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i, c being Integer for X, Y being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i, c being Integer for X, Y being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i, c being Integer for X, Y being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 be Int_position; ::_thesis: for i, c being Integer for X, Y being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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, c be Integer; ::_thesis: for X, Y being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 X, Y be set ; ::_thesis: ( ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 b = DataLoc ((s . a),i); defpred S1[ State of SCMPDS] means ( ( for x being Int_position st x in X holds $1 . x >= c + ($1 . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds $1 . x = s . x ) ); consider f being Function of (product (the_Values_of SCMPDS)),NAT such that A2: for s being State of SCMPDS holds ( ( s . (DataLoc ((s . a),i)) <= 0 implies f . s = 0 ) & ( s . (DataLoc ((s . a),i)) > 0 implies f . s = s . (DataLoc ((s . a),i)) ) ) by Th5; deffunc H1( State of SCMPDS) -> Element of NAT = f . $1; A5: for t being 0 -started State of SCMPDS st S1[t] & H1(t) = 0 holds t . (DataLoc ((s . a),i)) <= 0 by A2; assume A6: for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ) 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)))) ) ) ) A7: S1[s] by A6; assume A8: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in Y holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_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)))) ) ) A9: 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 A10: S1[t] and A11: t . a = s . a and A12: 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; consider v being State of SCMPDS such that A13: v = t and A14: for x being Int_position st x in X holds v . x >= c + (v . (DataLoc ((s . a),i))) and A15: for x being Int_position st x in Y holds v . x = s . x by A10; thus ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q ) by A8, A11, A12, A13, A15, A14; ::_thesis: ( H1( Initialize (IExec (I,Q,t))) < H1(t) & S1[ Initialize (IExec (I,Q,t))] ) thus H1( Initialize (IExec (I,Q,t))) < H1(t) ::_thesis: S1[ Initialize (IExec (I,Q,t))] proof A18: H1(t) = t . (DataLoc ((s . a),i)) by A2, A12; assume A19: H1( Initialize (IExec (I,Q,t))) >= H1(t) ; ::_thesis: contradiction H1( Initialize (IExec (I,Q,t))) = (Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i)) by A2, A12, A18, A19 .= (IExec (I,Q,t)) . (DataLoc ((s . a),i)) by SCMPDS_5:15 ; hence contradiction by A8, A11, A12, A10, A19, A18; ::_thesis: verum end; thus S1[ Initialize (IExec (I,Q,t))] ::_thesis: verum proof set v = Initialize (IExec (I,Q,t)); hereby ::_thesis: for x being Int_position st x in Y holds (Initialize (IExec (I,Q,t))) . x = s . x let x be Int_position; ::_thesis: ( x in X implies (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i))) ) assume x in X ; ::_thesis: (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i))) then (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) by A8, A11, A12, A10; then (Initialize (IExec (I,Q,t))) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) by SCMPDS_5:15; hence (Initialize (IExec (I,Q,t))) . x >= c + ((Initialize (IExec (I,Q,t))) . (DataLoc ((s . a),i))) by SCMPDS_5:15; ::_thesis: verum end; hereby ::_thesis: verum let x be Int_position; ::_thesis: ( x in Y implies (Initialize (IExec (I,Q,t))) . x = s . x ) assume A20: x in Y ; ::_thesis: (Initialize (IExec (I,Q,t))) . x = s . x then (IExec (I,Q,t)) . x = t . x by A8, A11, A12, A10; then (Initialize (IExec (I,Q,t))) . x = t . x by SCMPDS_5:15; hence (Initialize (IExec (I,Q,t))) . x = s . x by A13, A15, A20; ::_thesis: verum end; end; end; ( 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(A5, A7, A9); 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 A21: 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 SCMPDS_8:sch_4(A21, A5, A7, A9); hence IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) ; ::_thesis: verum end; theorem :: SCMPDS_8:28 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 being Int_position for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 be Int_position; ::_thesis: for i being Integer for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 Integer; ::_thesis: for X being set st ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 X be set ; ::_thesis: ( ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ) 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 b = DataLoc ((s . a),i); assume A2: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; ::_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)))) ) ) then A3: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {} holds t . x >= 0 + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in {} holds (IExec (I,Q,t)) . x >= 0 + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) ; A4: for x being Int_position st x in {} holds s . x >= 0 + (s . (DataLoc ((s . a),i))) ; for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in {} holds t . x >= 0 + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in {} holds (IExec (I,Q,t)) . x >= 0 + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x = t . x ) ) by A2; hence ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) by A4, Th27; ::_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 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)))) hence IExec ((while>0 (a,i,I)),P,s) = IExec ((while>0 (a,i,I)),P,(Initialize (IExec (I,P,s)))) by A4, A3, Th27; ::_thesis: verum end; theorem :: SCMPDS_8:29 for P being Instruction-Sequence of SCMPDS for s being 0 -started State of SCMPDS for I being halt-free shiftable Program of SCMPDS for a being Int_position for i, c being Integer for X being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 being Int_position for i, c being Integer for X being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 being Int_position for i, c being Integer for X being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 being Int_position for i, c being Integer for X being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 be Int_position; ::_thesis: for i, c being Integer for X being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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, c be Integer; ::_thesis: for X being set st ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 X be set ; ::_thesis: ( ( for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ) & ( for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 b = DataLoc ((s . a),i); assume A2: for x being Int_position st x in X holds s . x >= c + (s . (DataLoc ((s . a),i))) ; ::_thesis: ( ex t being 0 -started State of SCMPDS ex Q being Instruction-Sequence of SCMPDS st ( ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & t . a = s . a & t . (DataLoc ((s . a),i)) > 0 & not ( (IExec (I,Q,t)) . a = t . a & I is_closed_on t,Q & I is_halting_on t,Q & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ) 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 A3: for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) ) ; ::_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)))) ) ) then for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in {} holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in {} holds (IExec (I,Q,t)) . x = t . x ) ) ; hence ( while>0 (a,i,I) is_closed_on s,P & while>0 (a,i,I) is_halting_on s,P ) by A2, Th27; ::_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)))) ) for t being 0 -started State of SCMPDS for Q being Instruction-Sequence of SCMPDS st ( for x being Int_position st x in X holds t . x >= c + (t . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in {} holds t . x = s . x ) & 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 & (IExec (I,Q,t)) . (DataLoc ((s . a),i)) < t . (DataLoc ((s . a),i)) & ( for x being Int_position st x in X holds (IExec (I,Q,t)) . x >= c + ((IExec (I,Q,t)) . (DataLoc ((s . a),i))) ) & ( for x being Int_position st x in {} holds (IExec (I,Q,t)) . x = t . x ) ) by A3; hence ( 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)))) ) by A2, Th27; ::_thesis: verum end;