Copyright (c) 1999 Association of Mizar Users
environ vocabulary AMI_3, SCMPDS_2, AMI_1, INT_1, SCMPDS_4, AMI_2, RELAT_1, BOOLE, SCM_1, FUNCT_1, FUNCT_7, SCMFSA6A, FUNCT_4, SCMPDS_3, CARD_1, SCMFSA_7, ABSVALUE, ARYTM_1, RELOC, SCMFSA6B, SCMPDS_5, AMI_5, SCMFSA8A, CAT_1, UNIALG_2, SCMFSA7B, SCMFSA8B, SCMPDS_6; notation TARSKI, XBOOLE_0, XCMPLX_0, XREAL_0, ABSVALUE, RELAT_1, FUNCT_1, FUNCT_4, INT_1, NAT_1, STRUCT_0, AMI_1, AMI_2, AMI_3, AMI_5, FUNCT_7, SCMPDS_2, SCMPDS_3, CARD_1, SCMPDS_4, SCM_1, SCMPDS_5; constructors NAT_1, AMI_5, SCMFSA_4, SCMPDS_4, SCM_1, SCMPDS_5; clusters AMI_1, INT_1, FUNCT_1, SCMPDS_2, SCMFSA_4, SCMPDS_4, SCMPDS_5, NAT_1, FRAENKEL, XREAL_0, ORDINAL2, NUMBERS; requirements NUMERALS, REAL, SUBSET, BOOLE, ARITHM; theorems AMI_1, AMI_3, NAT_1, REAL_1, CQC_LANG, TARSKI, FUNCT_4, FUNCT_1, INT_1, AMI_5, SCMPDS_2, SCMPDS_3, ABSVALUE, SCMFSA6A, GRFUNC_1, AXIOMS, SCM_1, SCMFSA6B, SCMPDS_4, CQC_THE1, SCMFSA8A, SCMPDS_5, RELAT_1, XBOOLE_0, XBOOLE_1, XCMPLX_1; schemes NAT_1; begin :: Preliminaries reserve m,n for Nat, a for Int_position, i,j for Instruction of SCMPDS, s,s1,s2 for State of SCMPDS, k1 for Integer, loc for Instruction-Location of SCMPDS, I,J,K for Program-block; set A = the Instruction-Locations of SCMPDS; set D = SCM-Data-Loc; theorem Th1: :: S8A_Th3 for s being State of SCMPDS holds dom (s | the Instruction-Locations of SCMPDS) = the Instruction-Locations of SCMPDS proof let s be State of SCMPDS; thus dom (s | A) = dom s /\ A by RELAT_1:90 .= (D \/ {IC SCMPDS} \/ A) /\ A by SCMPDS_4:19 .= A by XBOOLE_1:21; end; theorem Th2: ::S8A_Th4 for s being State of SCMPDS st s is halting for k being Nat st LifeSpan s <= k holds CurInstr (Computation s).k = halt SCMPDS proof let s be State of SCMPDS; assume A1: s is halting; let k be Nat; assume A2: LifeSpan s <= k; CurInstr (Computation s).LifeSpan s = halt SCMPDS by A1,SCM_1:def 2; hence thesis by A2,AMI_1:52; end; theorem Th3: ::S8A_Th5 for s being State of SCMPDS st s is halting for k being Nat st LifeSpan s <= k holds IC (Computation s).k = IC (Computation s).LifeSpan s proof let s be State of SCMPDS; assume A1: s is halting; let k be Nat; assume A2: LifeSpan s <= k; defpred P[Nat] means LifeSpan s <= $1 implies IC (Computation s).$1 = IC (Computation s).LifeSpan s; A3: P[0] proof assume A4: LifeSpan s <= 0; 0 <= LifeSpan s by NAT_1:18; hence IC (Computation s).0 = IC (Computation s).LifeSpan s by A4,AXIOMS:21; end; A5: now let k be Nat; assume A6: P[k]; now assume A7: LifeSpan s <= k + 1; per cases by A7,REAL_1:def 5; suppose k + 1 = LifeSpan s; hence IC (Computation s).(k + 1) = IC (Computation s).LifeSpan s; suppose A8: k + 1 > LifeSpan s; then LifeSpan s <= k by NAT_1:38; then A9: CurInstr (Computation s).k = halt SCMPDS by A1,Th2; thus IC (Computation s).(k + 1) = IC Following (Computation s).k by AMI_1:def 19 .= IC Exec(halt SCMPDS,(Computation s).k) by A9,AMI_1:def 18 .= IC (Computation s).LifeSpan s by A6,A8,AMI_1:def 8,NAT_1:38; end; hence P[k + 1]; end; for k being Nat holds P[k] from Ind(A3,A5); hence thesis by A2; end; theorem Th4: ::S8A_Th6 for s1,s2 being State of SCMPDS holds s1,s2 equal_outside the Instruction-Locations of SCMPDS iff IC s1 = IC s2 & s1 | SCM-Data-Loc = s2 | SCM-Data-Loc proof let s1,s2 be State of SCMPDS; thus s1,s2 equal_outside the Instruction-Locations of SCMPDS implies IC s1 = IC s2 & s1 | D = s2 | D by SCMFSA6A:29,SCMPDS_4:24; assume A1: IC s1 = IC s2 & s1 | D = s2 | D; then for a being Int_position holds s1.a = s2.a by SCMPDS_4:23; hence s1,s2 equal_outside the Instruction-Locations of SCMPDS by A1,SCMPDS_4:11; end; theorem ::S8A_TI8 for s being State of SCMPDS, I being Program-block holds Initialized s +* Initialized I = s +* Initialized I proof let s be State of SCMPDS,I be Program-block; set SA0=Start-At inspos 0; A1: dom I misses dom SA0 by SCMPDS_4:54; thus Initialized s +* Initialized I = s +* SA0 +* Initialized I by SCMPDS_5:def 4 .= s +* SA0 +* (I +* SA0) by SCMPDS_4:def 2 .= s +* SA0 +* I +* SA0 by FUNCT_4:15 .= s +* (SA0 +* I) +* SA0 by FUNCT_4:15 .= s +* (I +* SA0) +* SA0 by A1,FUNCT_4:36 .= s +* I +* SA0 +* SA0 by FUNCT_4:15 .= s +* I +* (SA0 +* SA0) by FUNCT_4:15 .= s +* (I +* SA0) by FUNCT_4:15 .= s +* Initialized I by SCMPDS_4:def 2; end; theorem ::S8A_Th9 for I being Program-block, l being Instruction-Location of SCMPDS holds I c= I +* Start-At l proof let I be Program-block,l be Instruction-Location of SCMPDS; consider n such that A1: l = inspos n by SCMPDS_3:32; dom I misses dom Start-At l by A1,SCMPDS_4:54; hence I c= I +* Start-At l by FUNCT_4:33; end; theorem Th7: for s being State of SCMPDS, l being Instruction-Location of SCMPDS holds s | SCM-Data-Loc = (s +* Start-At l) | SCM-Data-Loc proof let s be State of SCMPDS; let l be Instruction-Location of SCMPDS; now let x be set; assume x in dom Start-At l; then x in {IC SCMPDS} by AMI_3:34; hence not x in D by SCMPDS_3:6,TARSKI:def 1; end; then dom Start-At l misses D by XBOOLE_0:3; hence s | D = (s +* Start-At l) | D by SCMFSA8A:2; end; theorem Th8: ::S8A_11 for s being State of SCMPDS,I being Program-block, l being Instruction-Location of SCMPDS holds s | SCM-Data-Loc = (s +* (I +* Start-At l)) | SCM-Data-Loc proof let s be State of SCMPDS,I be Program-block; let l be Instruction-Location of SCMPDS; now let x be set; assume x in dom (I +* Start-At l); then x in dom I \/ dom Start-At l by FUNCT_4:def 1; then x in dom I or x in dom Start-At l by XBOOLE_0:def 2; then A1: x in dom I or x in {IC SCMPDS} by AMI_3:34; per cases by A1,TARSKI:def 1; suppose A2: x in dom I; dom I c= the Instruction-Locations of SCMPDS by AMI_3:def 13; hence not x in D by A2,SCMPDS_2:10,XBOOLE_0:3; suppose x = IC SCMPDS; hence not x in D by SCMPDS_3:6; end; then dom (I +* Start-At l) misses D by XBOOLE_0:3; hence s | D = (s +* (I +* Start-At l)) | D by SCMFSA8A:2; end; theorem Th9: for s being State of SCMPDS,I being Program-block holds s | SCM-Data-Loc = (s +* Initialized I) | SCM-Data-Loc proof let s be State of SCMPDS,I be Program-block; Initialized I =I +* Start-At(inspos 0) by SCMPDS_4:def 2; hence thesis by Th8; end; theorem Th10: :: S8A_Th12 for s being State of SCMPDS, l being Instruction-Location of SCMPDS holds dom (s | the Instruction-Locations of SCMPDS) misses dom Start-At l proof let s be State of SCMPDS,l be Instruction-Location of SCMPDS; now let x be set; assume x in dom (s | A); then x is Instruction-Location of SCMPDS by Th1; then x <> IC SCMPDS by AMI_1:48; then not x in {IC SCMPDS} by TARSKI:def 1; hence not x in dom Start-At l by AMI_3:34; end; hence dom (s | A) misses dom Start-At l by XBOOLE_0:3; end; theorem ::S8A_Th14 for s being State of SCMPDS, I,J being Program-block, l being Instruction-Location of SCMPDS holds s +* (I +* Start-At l), s +* (J +* Start-At l) equal_outside the Instruction-Locations of SCMPDS proof let s be State of SCMPDS,I,J be Program-block; let l be Instruction-Location of SCMPDS; A1: IC (s +* (J +* Start-At l)) = IC (s +* J +* Start-At l) by FUNCT_4:15 .= l by AMI_5:79 .= IC (s +* I +* Start-At l) by AMI_5:79 .= IC (s +* (I +* Start-At l)) by FUNCT_4:15; now let a; A2: a in dom s & not a in dom (I +* Start-At l) & not a in dom (J +* Start-At l) by SCMPDS_2:49,SCMPDS_4:61; hence (s +* (J +* Start-At l)).a = s.a by FUNCT_4:12 .= (s +* (I +* Start-At l)).a by A2,FUNCT_4:12; end; hence thesis by A1,SCMPDS_4:11; end; theorem Th12: ::S8B_7 for s1,s2 be State of SCMPDS,I,J be Program-block holds s1 | SCM-Data-Loc = s2 | SCM-Data-Loc implies s1 +* Initialized I, s2 +* Initialized J equal_outside the Instruction-Locations of SCMPDS proof let s1,s2 be State of SCMPDS,I,J be Program-block; assume A1: s1 | D = s2 | D; set II=Initialized I, S1 = s1 +* II, IJ=Initialized J, S2 = s2 +* IJ; A2: S1 | D = s1 | D by Th9 .= S2 | D by A1,Th9; II c= S1 & IJ c= S2 by FUNCT_4:26; then IC S1 = inspos 0 & IC S2 = inspos 0 by SCMPDS_5:18; hence thesis by A2,Th4; end; theorem for I being programmed FinPartState of SCMPDS, x being set holds x in dom I implies I.x is Instruction of SCMPDS by SCMFSA8A:48; theorem Th14: ::S8B_Th4 for s being State of SCMPDS, l1,l2 being Instruction-Location of SCMPDS holds s +* Start-At l1 +* Start-At l2 = s +* Start-At l2 proof let s be State of SCMPDS; let l1,l2 be Instruction-Location of SCMPDS; A1: dom Start-At l1 = {IC SCMPDS} by AMI_3:34 .= dom Start-At l2 by AMI_3:34; thus s +* Start-At l1 +* Start-At l2 = s +* (Start-At l1 +* Start-At l2) by FUNCT_4:15 .= s +* Start-At l2 by A1,FUNCT_4:20; end; theorem Th15: card (i ';' I)= card I + 1 proof thus card (i ';' I) = card (Load i ';' I) by SCMPDS_4:def 4 .=card (Load i) + card I by SCMPDS_4:45 .=card I+1 by SCMPDS_5:6; end; theorem Th16: (i ';' I).inspos 0=i proof A1: i ';' I=Load i ';' I by SCMPDS_4:def 4; inspos 0 in dom Load i by SCMPDS_5:2; hence (i ';' I).inspos 0 =(Load i).inspos 0 by A1,SCMPDS_4:37 .=i by SCMPDS_5:4; end; theorem Th17: ::SP_15 I c= Initialized stop I proof set pI=stop I; pI= I ';' SCMPDS-Stop by SCMPDS_4:def 7; then A1: I c= pI by SCMPDS_4:40; pI c= Initialized pI by SCMPDS_4:9; hence thesis by A1,XBOOLE_1:1; end; theorem Th18: loc in dom I implies loc in dom (stop I) proof assume A1: loc in dom I; dom I c= dom (I ';' SCMPDS-Stop) by SCMPDS_4:39; then dom I c= dom stop I by SCMPDS_4:def 7; hence thesis by A1; end; theorem Th19: loc in dom I implies (stop I).loc=I.loc proof assume A1: loc in dom I; thus (stop I).loc =(I ';' SCMPDS-Stop).loc by SCMPDS_4:def 7 .=I.loc by A1,SCMPDS_4:37; end; theorem Th20: loc in dom I implies (Initialized stop I).loc=I.loc proof assume A1: loc in dom I; then loc in dom stop (I) by Th18; hence (Initialized stop I).loc=(stop I).loc by SCMPDS_4:33 .=I.loc by A1,Th19; end; theorem Th21: IC (s+* Initialized I)=inspos 0 proof Initialized I c= s+*Initialized I by FUNCT_4:26; hence thesis by SCMPDS_5:18; end; theorem Th22: CurInstr (s+* Initialized stop(i ';' I)) = i proof set iI=i ';' I, IsiI=Initialized stop iI, s3=s+*IsiI; card iI=card I +1 by Th15; then 0 < card iI by NAT_1:19; then A1: inspos 0 in dom iI by SCMPDS_4:1; iI c= IsiI by Th17; then A2: dom iI c= dom IsiI by GRFUNC_1:8; A3: IC s3 = inspos 0 by Th21; A4: inspos 0 in dom Load i by SCMPDS_5:2; s3.inspos 0 = IsiI.inspos 0 by A1,A2,FUNCT_4:14 .= iI.inspos 0 by A1,Th20 .=(Load i ';' I).inspos 0 by SCMPDS_4:def 4 .=(Load i).inspos 0 by A4,SCMPDS_4:37 .=i by SCMPDS_5:4; hence CurInstr s3 = i by A3,AMI_1:def 17; end; theorem Th23: for s being State of SCMPDS,m1,m2 being Nat st IC s=inspos m1 holds ICplusConst(s,m2)=inspos (m1+m2) proof let s be State of SCMPDS,m1,m2 be Nat; assume A1: IC s=inspos m1; consider m such that A2: m = IC s & ICplusConst(s,m2) = abs(m-2+2*m2)+2 by SCMPDS_2:def 20; A3: m=il.m1 by A1,A2,SCMPDS_3:def 2 .=2*m1 +2 by AMI_3:def 20; A4: 2*m1+2*m2 >= 0 by NAT_1:18; thus ICplusConst(s,m2) = abs(2*m1+2*m2)+2 by A2,A3,XCMPLX_1:26 .=2*m1+2*m2+2 by A4,ABSVALUE:def 1 .=2*(m1+m2)+2 by XCMPLX_1:8 .=il.(m1+m2) by AMI_3:def 20 .=inspos (m1+m2) by SCMPDS_3:def 2; end; theorem Th24: for I,J being Program-block holds Shift(stop J,card I) c= stop(I ';' J) proof let I,J be Program-block; stop(I ';' J) =I ';' J ';' SCMPDS-Stop by SCMPDS_4:def 7 .=I ';' (J ';' SCMPDS-Stop) by SCMPDS_4:46 .=I ';' (stop J) by SCMPDS_4:def 7; then stop(I ';' J) = I +* Shift(stop J, card I) by SCMPDS_4:def 3; hence thesis by FUNCT_4:26; end; theorem Th25: inspos(card I) in dom (stop I) & (stop I).inspos(card I) = halt SCMPDS proof set pI=stop I; card pI=card I+1 by SCMPDS_5:7; then card I <card pI by REAL_1:69; hence inspos(card I) in dom pI by SCMPDS_4:1; set SS=SCMPDS-Stop; A1: pI=I ';' SS by SCMPDS_4:def 7; pI.inspos(0+card I) =pI.(inspos 0+card I) by SCMPDS_3:def 3 .=halt SCMPDS by A1,SCMPDS_4:38,73; hence thesis; end; theorem Th26: for x,l being Instruction-Location of SCMPDS holds IExec(J,s).x = (IExec(I,s) +* Start-At l).x proof let x,l be Instruction-Location of SCMPDS; A1: dom Start-At l = {IC SCMPDS} by AMI_3:34; x <> IC SCMPDS by AMI_1:48; then A2: not x in dom Start-At l by A1,TARSKI:def 1; A3: dom (s | A) = A by Th1; thus IExec(J,s).x = (Result(s+*Initialized stop J) +* s | A).x by SCMPDS_4:def 8 .=(s | A).x by A3,FUNCT_4:14 .= (Result(s+*Initialized stop I) +* s | A).x by A3,FUNCT_4:14 .= IExec(I,s).x by SCMPDS_4:def 8 .= (IExec(I,s) +* Start-At l).x by A2,FUNCT_4:12; end; theorem Th27: for x,l being Instruction-Location of SCMPDS holds IExec(I,s).x = (s +* Start-At l).x proof let x,l be Instruction-Location of SCMPDS; A1: dom Start-At l = {IC SCMPDS} by AMI_3:34; x <> IC SCMPDS by AMI_1:48; then A2: not x in dom Start-At l by A1,TARSKI:def 1; A3: dom (s | A) = A by Th1; thus IExec(I,s).x = (Result(s+*Initialized stop I) +* s | A).x by SCMPDS_4:def 8 .=(s | A).x by A3,FUNCT_4:14 .= s.x by FUNCT_1:72 .= (s+* Start-At l).x by A2,FUNCT_4:12; end; theorem ::S8B_12 for s being State of SCMPDS, i being No-StopCode parahalting Instruction of SCMPDS,J being parahalting shiftable Program-block,a being Int_position holds IExec(i ';' J,s).a = IExec(J,Exec(i,Initialized s)).a proof let s be State of SCMPDS,i be No-StopCode parahalting Instruction of SCMPDS,J be parahalting shiftable Program-block, a be Int_position; thus IExec(i ';' J,s).a = IExec(Load i ';' J,s).a by SCMPDS_4:def 4 .= IExec(J,IExec(Load i,s)).a by SCMPDS_5:39 .= IExec(J,Exec(i,Initialized s)).a by SCMPDS_5:45; end; theorem Th29: for a being Int_position,k1,k2 being Integer holds (a,k1)<>0_goto k2 <> halt SCMPDS proof let a be Int_position,k1,k2 be Integer; InsCode ((a,k1)<>0_goto k2) = 4 by SCMPDS_2:25; hence thesis by SCMPDS_2:21,93; end; theorem Th30: for a being Int_position,k1,k2 being Integer holds (a,k1)<=0_goto k2 <> halt SCMPDS proof let a be Int_position,k1,k2 be Integer; InsCode ((a,k1)<=0_goto k2) = 5 by SCMPDS_2:26; hence thesis by SCMPDS_2:21,93; end; theorem Th31: for a being Int_position,k1,k2 being Integer holds (a,k1)>=0_goto k2 <> halt SCMPDS proof let a be Int_position,k1,k2 be Integer; InsCode ((a,k1)>=0_goto k2) = 6 by SCMPDS_2:27; hence thesis by SCMPDS_2:21,93; end; definition let k1; func Goto k1 -> Program-block equals :Def1: Load (goto k1); coherence; end; definition let n be Nat; cluster goto (n+1) -> No-StopCode; correctness by SCMPDS_5:25; cluster goto -(n+1) -> No-StopCode; correctness proof -(n+1) <> 0 by XCMPLX_1:135; hence thesis by SCMPDS_5:25; end; end; definition let n be Nat; cluster Goto (n+1) -> No-StopCode; correctness proof Goto (n+1) =Load goto (n+1) by Def1; hence thesis; end; cluster Goto -(n+1) -> No-StopCode; correctness proof Goto -(n+1) =Load goto -(n+1) by Def1; hence thesis; end; end; theorem Th32: card Goto k1 = 1 proof thus card Goto k1 = card Load (goto k1) by Def1 .=1 by SCMPDS_5:6; end; Lm1: inspos 0 in dom (inspos 0 .--> goto k1) & (inspos 0 .--> goto k1).inspos 0 = goto k1 proof dom (inspos 0 .--> goto k1) = {inspos 0} by CQC_LANG:5; hence inspos 0 in dom (inspos 0 .--> goto k1) by TARSKI:def 1; thus (inspos 0 .--> goto k1).inspos 0 = goto k1 by CQC_LANG:6; end; theorem Th33: ::S8A_Th47 inspos 0 in dom Goto k1 & (Goto k1).inspos 0 = goto k1 proof Goto k1 = Load (goto k1) by Def1 .=(inspos 0 .--> goto k1) by SCMPDS_4:def 1; hence thesis by Lm1; end; begin :: The predicates of is_closed_on and is_halting_on definition let I be Program-block; let s be State of SCMPDS; pred I is_closed_on s means :Def2: for k being Nat holds IC (Computation (s +* Initialized stop I )).k in dom stop I; pred I is_halting_on s means :Def3: s +* Initialized stop I is halting; end; theorem Th34: ::S7B_Th24 for I being Program-block holds I is paraclosed iff for s being State of SCMPDS holds I is_closed_on s proof let I be Program-block; set IsI=Initialized stop(I); hereby assume A1: I is paraclosed; let s be State of SCMPDS; IsI c= s +* IsI by FUNCT_4:26; then for n holds IC (Computation (s +* IsI)).n in dom stop I by A1,SCMPDS_4:def 9; hence I is_closed_on s by Def2; end; assume A2: for s being State of SCMPDS holds I is_closed_on s; now let s be State of SCMPDS; let k be Nat; assume IsI c= s; then I is_closed_on s & s = s +* IsI by A2,AMI_5:10; hence IC (Computation s).k in dom stop I by Def2; end; hence I is paraclosed by SCMPDS_4:def 9; end; theorem Th35: ::S7B_25 for I being Program-block holds I is parahalting iff for s being State of SCMPDS holds I is_halting_on s proof let I be Program-block; set IsI=Initialized stop(I); hereby assume A1: I is parahalting; let s be State of SCMPDS; IsI c= s +* IsI by FUNCT_4:26; then s +* IsI is halting by A1,SCMPDS_4:63; hence I is_halting_on s by Def3; end; assume A2: for s being State of SCMPDS holds I is_halting_on s; now let s be State of SCMPDS; assume IsI c= s; then I is_halting_on s & s = s +* IsI by A2,AMI_5:10; hence s is halting by Def3; end; then IsI is halting by AMI_1:def 26; hence I is parahalting by SCMPDS_4:def 10; end; theorem Th36: for s1,s2 being State of SCMPDS, I being Program-block st s1 | SCM-Data-Loc = s2 | SCM-Data-Loc holds I is_closed_on s1 implies I is_closed_on s2 proof let s1,s2 be State of SCMPDS,I be Program-block; set pI=stop I, IsI=Initialized pI, C1 = Computation (s1 +* IsI), C2 = Computation (s2 +* IsI); assume A1: s1 | D = s2 | D; assume A2: I is_closed_on s1; defpred P[Nat] means IC C1.$1 = IC C2.$1 & CurInstr C1.$1 = CurInstr C2.$1 & C1.$1 | D = C2.$1 | D; A3: IC SCMPDS in dom IsI by SCMPDS_4:7; A4: C1.0 = s1 +* IsI & C2.0 = s2 +* IsI by AMI_1:def 19; A5: IC C1.0 = C1.0.IC SCMPDS by AMI_1:def 15 .= IsI.IC SCMPDS by A3,A4,FUNCT_4:14 .= inspos 0 by SCMPDS_4:29; A6: IC C2.0 = C2.0.IC SCMPDS by AMI_1:def 15 .= IsI.IC SCMPDS by A3,A4,FUNCT_4:14 .= inspos 0 by SCMPDS_4:29; A7: pI c= IsI by SCMPDS_4:9; then A8: dom pI c= dom IsI by GRFUNC_1:8; A9: inspos 0 in dom pI by SCMPDS_4:75; A10: CurInstr C1.0 = C1.0.IC C1.0 by AMI_1:def 17 .= IsI.inspos 0 by A4,A5,A8,A9,FUNCT_4:14 .= C2.0.IC C2.0 by A4,A6,A8,A9,FUNCT_4:14 .= CurInstr C2.0 by AMI_1:def 17; C1.0 | D = s1 | D by A4,Th9 .= C2.0 | D by A1,A4,Th9; then A11: P[0] by A5,A6,A10; A12: now let k be Nat; assume A13: P[k]; then for a holds C1.k.a = C2.k.a by SCMPDS_4:23; then C1.k,C2.k equal_outside A by A13,SCMPDS_4:11; then A14: Exec(CurInstr C1.k,C1.k),Exec(CurInstr C1.k,C2.k) equal_outside A by SCMPDS_4:15; thus P[k+1] proof IsI c= s1 +* IsI & IsI c= s2 +* IsI by FUNCT_4:26; then A15: pI c= s1 +* IsI & pI c= s2 +* IsI by A7,XBOOLE_1:1; then A16: pI c= C1.k & pI c= C2.k by AMI_3:38; A17: pI c= C1.(k + 1) & pI c= C2.(k + 1) by A15,AMI_3:38; A18: IC C1.k in dom pI by A2,Def2; A19: CurInstr C1.k = C1.k.IC C1.k by AMI_1:def 17 .= pI.IC C1.k by A16,A18,GRFUNC_1:8 .= C2.k.IC C2.k by A13,A16,A18,GRFUNC_1:8 .= CurInstr C2.k by AMI_1:def 17; A20: C1.(k + 1) = Following C1.k by AMI_1:def 19 .= Exec(CurInstr C1.k,C1.k) by AMI_1:def 18; A21: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18; hence A22: IC C1.(k + 1) = IC C2.(k + 1) by A14,A19,A20,SCMFSA6A:29; A23: IC C1.(k + 1) in dom pI by A2,Def2; thus CurInstr C1.(k + 1) = C1.(k + 1).IC C1.(k + 1) by AMI_1:def 17 .= pI.IC C1.(k + 1) by A17,A23,GRFUNC_1:8 .= C2.(k + 1).IC C2.(k + 1) by A17,A22,A23,GRFUNC_1:8 .= CurInstr C2.(k + 1) by AMI_1:def 17; thus C1.(k + 1) | D = C2.(k + 1) | D by A14,A19,A20,A21,SCMPDS_4:24; end; end; now let k be Nat; A24: IC C1.k in dom pI by A2,Def2; for k being Nat holds P[k] from Ind(A11,A12); hence IC C2.k in dom pI by A24; end; hence I is_closed_on s2 by Def2; end; theorem ::S8B_Th8 for s1,s2 being State of SCMPDS,I being Program-block st s1 | SCM-Data-Loc = s2 | SCM-Data-Loc holds I is_closed_on s1 & I is_halting_on s1 implies I is_closed_on s2 & I is_halting_on s2 proof let s1,s2 be State of SCMPDS,I be Program-block; set pI=stop I, IsI=Initialized pI, C1 = Computation (s1 +* IsI), C2 = Computation (s2 +* IsI); assume A1: s1 | D = s2 | D; assume A2: I is_closed_on s1; assume I is_halting_on s1; then s1 +* IsI is halting by Def3; then consider m such that A3: CurInstr C1.m = halt SCMPDS by AMI_1:def 20; defpred P[Nat] means IC C1.$1 = IC C2.$1 & CurInstr C1.$1 = CurInstr C2.$1 & C1.$1 | D = C2.$1 | D; A4: IC SCMPDS in dom IsI by SCMPDS_4:7; A5: C1.0 = s1 +* IsI & C2.0 = s2 +* IsI by AMI_1:def 19; A6: IC C1.0 = C1.0.IC SCMPDS by AMI_1:def 15 .= IsI.IC SCMPDS by A4,A5,FUNCT_4:14 .= inspos 0 by SCMPDS_4:29; A7: IC C2.0 = C2.0.IC SCMPDS by AMI_1:def 15 .= IsI.IC SCMPDS by A4,A5,FUNCT_4:14 .= inspos 0 by SCMPDS_4:29; A8: pI c= IsI by SCMPDS_4:9; then A9: dom pI c= dom IsI by GRFUNC_1:8; A10: inspos 0 in dom pI by SCMPDS_4:75; A11: CurInstr C1.0 = C1.0.IC C1.0 by AMI_1:def 17 .= IsI.inspos 0 by A5,A6,A9,A10,FUNCT_4:14 .= C2.0.IC C2.0 by A5,A7,A9,A10,FUNCT_4:14 .= CurInstr C2.0 by AMI_1:def 17; C1.0,C2.0 equal_outside A by A1,A5,Th12; then A12: P[0] by A6,A7,A11,SCMPDS_4:24; A13: now let k be Nat; assume A14: P[k]; then for a holds C1.k.a = C2.k.a by SCMPDS_4:23; then C1.k,C2.k equal_outside A by A14,SCMPDS_4:11; then A15: Exec(CurInstr C1.k,C1.k),Exec(CurInstr C1.k,C2.k) equal_outside A by SCMPDS_4:15; thus P[k+1] proof IsI c= s1 +* IsI & IsI c= s2 +* IsI by FUNCT_4:26; then A16: pI c= s1 +* IsI & pI c= s2 +* IsI by A8,XBOOLE_1:1; then A17: pI c= C1.k & pI c= C2.k by AMI_3:38; A18: pI c= C1.(k + 1) & pI c= C2.(k + 1) by A16,AMI_3:38; A19: IC C1.k in dom pI by A2,Def2; A20: CurInstr C1.k = C1.k.IC C1.k by AMI_1:def 17 .= pI.IC C1.k by A17,A19,GRFUNC_1:8 .= C2.k.IC C2.k by A14,A17,A19,GRFUNC_1:8 .= CurInstr C2.k by AMI_1:def 17; A21: C1.(k + 1) = Following C1.k by AMI_1:def 19 .= Exec(CurInstr C1.k,C1.k) by AMI_1:def 18; A22: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18; hence A23: IC C1.(k + 1) = IC C2.(k + 1) by A15,A20,A21,SCMFSA6A:29; A24: IC C1.(k + 1) in dom pI by A2,Def2; thus CurInstr C1.(k + 1) = C1.(k + 1).IC C1.(k + 1) by AMI_1:def 17 .= pI.IC C1.(k + 1) by A18,A24,GRFUNC_1:8 .= C2.(k + 1).IC C2.(k + 1) by A18,A23,A24,GRFUNC_1:8 .= CurInstr C2.(k + 1) by AMI_1:def 17; thus C1.(k + 1) | D = C2.(k + 1) | D by A15,A20,A21,A22,SCMPDS_4:24; end; end; for k being Nat holds P[k] from Ind(A12,A13); then CurInstr C2.m = halt SCMPDS by A3; then A25: s2 +* IsI is halting by AMI_1:def 20; now let k be Nat; A26: IC C1.k in dom pI by A2,Def2; for k being Nat holds P[k] from Ind(A12,A13); hence IC C2.k in dom pI by A26; end; hence I is_closed_on s2 by Def2; thus I is_halting_on s2 by A25,Def3; end; theorem Th38: ::S8B_Th9 for s being State of SCMPDS, I,J being Program-block holds I is_closed_on s iff I is_closed_on s +* Initialized J proof let s be State of SCMPDS,I,J be Program-block; s | D = (s +* Initialized J) | D by Th9; hence thesis by Th36; end; theorem Th39: for I,J being Program-block,s being State of SCMPDS st I is_closed_on s & I is_halting_on s holds (for k being Nat st k <= LifeSpan (s +* Initialized stop I) holds IC (Computation (s +* Initialized stop I)).k = IC (Computation (s +* Initialized stop (I ';' J))).k) & (Computation (s +* Initialized stop I)). (LifeSpan (s +* Initialized stop I)) | SCM-Data-Loc = (Computation (s +* Initialized stop (I ';' J))). (LifeSpan (s +* Initialized stop I)) | SCM-Data-Loc proof let I,J be Program-block,s be State of SCMPDS; assume A1: I is_closed_on s; assume A2: I is_halting_on s; set pI=stop I, IsI=Initialized pI, pIJ=stop (I ';' J), IsJ=Initialized pIJ, s1=s +* IsI, IL=the Instruction-Locations of SCMPDS; A3: s1 is halting by A2,Def3; A4: IsI=pI +* Start-At inspos 0 by SCMPDS_4:def 2; A5: IsI c= s1 by FUNCT_4:26; then A6: s1 +* IsJ = s1 +* pIJ by SCMPDS_4:34; defpred X[Nat] means $1 <= LifeSpan s1 implies (Computation s1).$1,(Computation (s1+*IsJ)).$1 equal_outside IL; (Computation s1).0 = s1 & (Computation (s1+*IsJ)).0 = s1+*IsJ by AMI_1:def 19; then A7: X[0] by A6,SCMFSA6A:27; A8: for m st X[m] holds X[m+1] proof let m; assume A9: m <= LifeSpan s1 implies (Computation s1).m,(Computation (s1+*IsJ)).m equal_outside IL; assume A10: m+1 <= LifeSpan s1; then A11: m < LifeSpan s1 by NAT_1:38; set Cs = Computation s1, CsIJ = Computation (s1+*IsJ); A12: Cs.(m+1) = Following Cs.m by AMI_1:def 19 .= Exec(CurInstr Cs.m,Cs.m) by AMI_1:def 18; A13: CsIJ.(m+1) = Following CsIJ.m by AMI_1:def 19 .= Exec(CurInstr CsIJ.m,CsIJ.m) by AMI_1:def 18; A14: IC (Cs.m) = IC (CsIJ.m) by A9,A10,NAT_1:38,SCMFSA6A:29; A15: IC Cs.m in dom pI by A1,Def2; dom pI misses dom Start-At inspos 0 by SCMPDS_4:54; then pI c= pI +* Start-At inspos 0 by FUNCT_4:33; then pI c= s1 by A4,A5,XBOOLE_1:1; then A16: pI c= Cs.m by AMI_3:38; pIJ c= s1+* IsJ by A6,FUNCT_4:26; then A17: pIJ c= CsIJ.m by AMI_3:38; A18: CurInstr(Cs.m) = (Cs.m).IC (Cs.m) by AMI_1:def 17 .= pI.IC (Cs.m) by A15,A16,GRFUNC_1:8; then pI.IC(Cs.m) <> halt SCMPDS by A3,A11,SCM_1:def 2; then A19: IC Cs.m in dom I by A15,SCMPDS_5:3; set JS=J ';' SCMPDS-Stop; A20: pIJ =I ';' J ';' SCMPDS-Stop by SCMPDS_4:def 7 .=I ';' JS by SCMPDS_4:46; dom(I ';' JS) = dom (I +* Shift(JS, card I)) by SCMPDS_4:def 3 .= dom I \/ dom Shift(JS, card I) by FUNCT_4:def 1; then A21: dom I c= dom(I ';' JS) by XBOOLE_1:7; CurInstr(Cs.m)= (I ';' SCMPDS-Stop).IC (Cs.m) by A18,SCMPDS_4:def 7 .=I.IC (Cs.m) by A19,SCMPDS_4:37 .=pIJ.IC(Cs.m) by A19,A20,SCMPDS_4:37 .=(CsIJ.m).IC(Cs.m) by A17,A19,A20,A21,GRFUNC_1:8 .= CurInstr(CsIJ.m) by A14,AMI_1:def 17; hence (Computation s1).(m+1),(Computation(s1+*IsJ)).(m+1) equal_outside IL by A9,A10,A12,A13,NAT_1:38,SCMPDS_4:15; end; A22: s1+*IsJ=s +* (IsI +* IsJ) by FUNCT_4:15 .=s+*IsJ by SCMPDS_5:17; A23: for m holds X[m] from Ind(A7,A8); hereby let k be Nat; assume k <= LifeSpan s1; then (Computation s1).k,(Computation (s1 +* IsJ)).k equal_outside IL by A23; hence IC (Computation s1).k =IC (Computation (s +* IsJ)).k by A22, SCMFSA6A:29; end; set m=LifeSpan s1; (Computation s1).m,(Computation (s1 +* IsJ)).m equal_outside IL by A23; hence (Computation s1).m | D =(Computation (s +* IsJ)).m | D by A22,SCMPDS_4: 24 ; end; theorem Th40: for I being Program-block,k be Nat st I is_closed_on s & I is_halting_on s & k < LifeSpan (s +* Initialized stop(I)) holds IC (Computation (s +* Initialized stop(I))).k in dom I proof let I be Program-block,k be Nat; set IsI=Initialized stop(I), ss= s +* IsI, m=LifeSpan ss, Sp=SCMPDS-Stop; assume A1: I is_closed_on s & I is_halting_on s & k < m; A2: stop I = I ';' Sp by SCMPDS_4:def 7; set Sk= (Computation ss).k, Ik=IC Sk; A3: IsI c= ss by FUNCT_4:26; A4: Ik in dom stop(I) by A1,Def2; A5: ss is halting by A1,Def3; stop I c= IsI by SCMPDS_4:9; then A6: stop I c= ss by A3,XBOOLE_1:1; consider n such that A7: inspos n= Ik by SCMPDS_3:32; card stop I=card I + 1 by SCMPDS_5:7; then n < card I + 1 by A4,A7,SCMPDS_4:1; then A8: n <= card I by INT_1:20; now assume A9: n = card I; CurInstr Sk = Sk.Ik by AMI_1:def 17 .=ss.Ik by AMI_1:54 .=(stop I).inspos(0+n) by A4,A6,A7,GRFUNC_1:8 .=(stop I).(inspos 0+n) by SCMPDS_3:def 3 .=halt SCMPDS by A2,A9,SCMPDS_4:38,73; hence contradiction by A1,A5,SCM_1:def 2; end; then n < card I by A8,REAL_1:def 5; hence thesis by A7,SCMPDS_4:1; end; theorem Th41: for I,J being Program-block,s being State of SCMPDS,k being Nat st I is_closed_on s & I is_halting_on s & k < LifeSpan (s +* Initialized stop I) holds CurInstr (Computation (s +* Initialized stop I)).k = CurInstr (Computation (s +* Initialized stop (I ';' J))).k proof let I,J be Program-block,s be State of SCMPDS,k be Nat; set IsI=Initialized stop I, IsJ=Initialized stop (I ';' J), s1=s+*IsI, s2=s+*IsJ; set s3=(Computation s1).k, s4=(Computation s2).k, SS=SCMPDS-Stop; assume A1: I is_closed_on s & I is_halting_on s & k < LifeSpan s1; then A2: IC s3 in dom I by Th40; A3: IC s3= IC s4 by A1,Th39; A4: IC s3 in dom stop(I) by A1,Def2; A5: IsI c= s1 by FUNCT_4:26; stop I c= IsI by SCMPDS_4:9; then A6: stop I c= s1 by A5,XBOOLE_1:1; A7: dom stop I c= dom stop (I ';' J) by SCMPDS_5:16; A8: IsJ c= s2 by FUNCT_4:26; stop (I ';' J) c= IsJ by SCMPDS_4:9; then A9: stop (I ';' J) c= s2 by A8,XBOOLE_1:1; A10: stop I = I ';' SS by SCMPDS_4:def 7; A11: stop (I ';' J) = I ';' J ';' SS by SCMPDS_4:def 7 .=I ';' (J ';' SS) by SCMPDS_4:46; thus CurInstr s3 =s3.IC s3 by AMI_1:def 17 .=s1.IC s3 by AMI_1:54 .=(stop I).IC s3 by A4,A6,GRFUNC_1:8 .=I.IC s3 by A2,A10,SCMPDS_4:37 .=(stop (I ';' J)).IC s3 by A2,A11,SCMPDS_4:37 .=s2.IC s4 by A3,A4,A7,A9,GRFUNC_1:8 .=s4.IC s4 by AMI_1:54 .=CurInstr s4 by AMI_1:def 17; end; theorem Th42: ::SCMPDS_5:32 for I being No-StopCode Program-block,s being State of SCMPDS, k being Nat st I is_closed_on s & I is_halting_on s & k < LifeSpan (s +* Initialized stop I) holds CurInstr (Computation (s +* Initialized stop I)).k <> halt SCMPDS proof let I be No-StopCode Program-block,s be State of SCMPDS,k be Nat; set IsI=Initialized stop(I), ss=s +* IsI, s2=(Computation ss).k; assume A1: I is_closed_on s & I is_halting_on s & k < LifeSpan ss; A2: IsI c= ss by FUNCT_4:26; I c= IsI by Th17; then I c= ss by A2,XBOOLE_1:1; then A3: I c= s2 by AMI_3:38; A4: IC s2 in dom I by A1,Th40; CurInstr s2=s2.IC s2 by AMI_1:def 17 .=I.IC s2 by A3,A4,GRFUNC_1:8; hence thesis by A4,SCMPDS_5:def 3; end; theorem Th43: for I being No-StopCode Program-block,s being State of SCMPDS st I is_closed_on s & I is_halting_on s holds IC (Computation (s +* Initialized stop I)). LifeSpan (s +* Initialized stop I) = inspos card I proof let I be No-StopCode Program-block,s be State of SCMPDS; set IsI=Initialized stop(I), s1=s +* IsI; assume A1: I is_closed_on s & I is_halting_on s; then A2: s1 is halting by Def3; A3: IsI c= s1 by FUNCT_4:26; I c= IsI by Th17; then A4: I c= s1 by A3,XBOOLE_1:1; set Css=(Computation s1).LifeSpan s1; A5: IC Css in dom stop(I) by A1,Def2; consider n such that A6: inspos n= IC Css by SCMPDS_3:32; now assume A7: IC Css in dom I; then I.IC Css=s1.IC Css by A4,GRFUNC_1:8 .=Css.IC Css by AMI_1:54 .=CurInstr Css by AMI_1:def 17 .=halt SCMPDS by A2,SCM_1:def 2; hence contradiction by A7,SCMPDS_5:def 3; end; then A8: n >= card I by A6,SCMPDS_4:1; card stop I =card I + 1 by SCMPDS_5:7; then n < card I + 1 by A5,A6,SCMPDS_4:1; then n <= card I by NAT_1:38; hence IC (Computation s1).LifeSpan s1 =inspos card I by A6,A8,AXIOMS:21; end; Lm2: for I being No-StopCode Program-block, J being Program-block, s being State of SCMPDS st I is_closed_on s & I is_halting_on s holds IC (Computation (s +* Initialized stop ( I ';' Goto (card J + 1) ';' J ))). (LifeSpan (s +* Initialized stop I) + 1) = inspos (card I + card J + 1) & (Computation (s +* Initialized stop I)). (LifeSpan (s +* Initialized stop I)) | D = (Computation (s +* Initialized stop (I ';' Goto (card J + 1) ';' J))). (LifeSpan (s +* Initialized stop I) + 1) | D & (for k being Nat st k <= LifeSpan (s +* Initialized stop I) holds CurInstr (Computation (s +*Initialized stop (I ';' Goto (card J + 1) ';' J ))).k <> halt SCMPDS) & IC (Computation (s +* Initialized stop (I ';' Goto (card J + 1) ';' J))). (LifeSpan (s +* Initialized stop I)) = inspos card I & s +* Initialized stop (I ';' Goto (card J + 1) ';' J) is halting & LifeSpan (s +* Initialized stop (I ';' Goto (card J + 1) ';' J)) = LifeSpan (s +* Initialized stop I) + 1 proof let I be No-StopCode Program-block, J be Program-block, s be State of SCMPDS; assume A1: I is_closed_on s; assume A2: I is_halting_on s; set G1=Goto (card J + 1), SS = SCMPDS-Stop, J2 = G1 ';' J ';' SS, IJ=I ';' G1 ';' J, pJ=stop IJ, IsJ=Initialized pJ, s1 = s +* Initialized stop I, s2 = s +* IsJ; A3: pJ =I ';' G1 ';' J ';' SS by SCMPDS_4:def 7 .=I ';' (G1 ';' J) ';' SS by SCMPDS_4:46 .=I ';' J2 by SCMPDS_4:46; A4: IJ =I ';' (G1 ';' J) by SCMPDS_4:46; A5: card (G1 ';' J) = card G1 + card J by SCMPDS_4:45 .=1 + card J by Th32; pJ c= IsJ by SCMPDS_4:9; then A6: dom pJ c= dom IsJ by GRFUNC_1:8; A7: card pJ = card I + card J2 by A3,SCMPDS_4:45; A8: card J2 = card (G1 ';' (J ';' SS)) by SCMPDS_4:46 .=card G1 + card (J ';' SS) by SCMPDS_4:45 .= 1 + card (J ';' SS) by Th32; then 0 + 1 <= card J2 by NAT_1:29; then A9: 0 < card J2 by NAT_1:38; then card I + 0 < card pJ by A7,REAL_1:53; then A10: inspos card I in dom pJ by SCMPDS_4:1; A11: inspos 0 in dom J2 by A9,SCMPDS_4:1; A12: inspos 0 in dom G1 by Th33; A13: J2.inspos 0 = (G1 ';' (J ';' SS)).inspos 0 by SCMPDS_4:46 .=G1.inspos 0 by A12,SCMPDS_4:37 .=goto (card J + 1) by Th33; set sm=(Computation s2).LifeSpan s1; A14: (Computation s1).(LifeSpan s1) | D = sm | D & IC (Computation s1).(LifeSpan s1) =IC sm by A1,A2,A4,Th39; then A15: IC sm = inspos card I by A1,A2,Th43; then A16: CurInstr (Computation s2).(LifeSpan s1) = (Computation s2).(LifeSpan s1).inspos card I by AMI_1:def 17 .= s2.inspos card I by AMI_1:54 .= IsJ.inspos card I by A6,A10,FUNCT_4:14 .= (I ';' J2).inspos(0+card I) by A3,A10,SCMPDS_4:33 .= (I ';' J2).(inspos 0 + card I) by SCMPDS_3:def 3 .= goto (card J + 1) by A11,A13,SCMPDS_4:38; A17: card J2 = 1 + (card J + card SS) by A8,SCMPDS_4:45 .= card J + (1 + card SS) by XCMPLX_1:1; then card J + 1 < card J2 by REAL_1:53,SCMPDS_4:74; then A18: inspos (card J + 1) in dom J2 by SCMPDS_4:1; card pJ = card I + card J + (1 + card SS) by A7,A17,XCMPLX_1:1 .= card I + card J + 1 + 1 by SCMPDS_4:74,XCMPLX_1:1; then card I + card J + 1 < card pJ by NAT_1:38; then A19: inspos (card I + card J + 1) in dom pJ by SCMPDS_4:1; A20: J2.inspos (card J + 1)=J2.inspos(0+card (G1 ';' J)) by A5 .=J2.(inspos 0+ card (G1 ';' J)) by SCMPDS_3:def 3 .=halt SCMPDS by SCMPDS_4:38,73; thus IC (Computation s2).(LifeSpan s1 + 1) = IC Following sm by AMI_1:def 19 .= IC Exec (goto (card J + 1),sm) by A16,AMI_1:def 18 .= Exec (goto (card J + 1),sm).IC SCMPDS by AMI_1:def 15 .= ICplusConst(sm,card J +1) by SCMPDS_2:66 .=inspos (card I + (card J + 1)) by A15,Th23 .=inspos (card I + card J + 1) by XCMPLX_1:1; then A21: CurInstr (Computation s2).(LifeSpan s1 + 1) = (Computation s2).(LifeSpan s1 + 1).inspos (card I + card J + 1) by AMI_1:def 17 .= s2.inspos (card I + card J + 1) by AMI_1:54 .= IsJ.inspos (card I + card J + 1) by A6,A19,FUNCT_4:14 .= (I ';' J2).inspos (card I + card J + 1) by A3,A19,SCMPDS_4:33 .= (I ';' J2).inspos(card I+(card J+1)) by XCMPLX_1:1 .= (I ';' J2).(inspos (card J+1)+card I) by SCMPDS_3:def 3 .= halt SCMPDS by A18,A20,SCMPDS_4:38; now let a; thus (Computation s2).(LifeSpan s1 + 1).a = (Following sm).a by AMI_1:def 19 .= Exec(goto (card J + 1),sm).a by A16,AMI_1:def 18 .= sm.a by SCMPDS_2:66; end; hence (Computation s1).(LifeSpan s1) | D = (Computation s2).(LifeSpan s1 + 1) | D by A14,SCMPDS_4:23; thus A22: now let k be Nat; assume A23: k <= LifeSpan s1; per cases; suppose A24: k < LifeSpan s1; then CurInstr (Computation s1).k <> halt SCMPDS by A1,A2,Th42; hence CurInstr (Computation s2).k <> halt SCMPDS by A1,A2,A4,A24,Th41; suppose LifeSpan s1 <= k; then k = LifeSpan s1 by A23,AXIOMS:21; hence CurInstr (Computation s2).k <> halt SCMPDS by A16,SCMPDS_2:85; end; A25: now let k be Nat; assume CurInstr (Computation s2).k = halt SCMPDS; then LifeSpan s1 < k by A22; hence LifeSpan s1+1 <= k by INT_1:20; end; thus IC (Computation s2).(LifeSpan s1) = inspos card I by A1,A2,A14,Th43; thus s2 is halting by A21,AMI_1:def 20; hence LifeSpan s2 = LifeSpan s1 + 1 by A21,A25,SCM_1:def 2; end; theorem Th44: ::S8A_58 for I,J being Program-block,s being State of SCMPDS st I is_closed_on s & I is_halting_on s holds I ';' Goto (card J + 1) ';' J is_halting_on s & I ';' Goto (card J + 1) ';' J is_closed_on s proof let I,J be Program-block,s be State of SCMPDS; assume A1: I is_closed_on s; assume A2: I is_halting_on s; set G = Goto (card J + 1), IJ = I ';' G ';' J, J2 = I ';' (G ';' J), pJ = stop J2, IsJ=Initialized pJ, pI =stop I, IsI=Initialized pI, s1 = s +* IsI, s2 = s +* IsJ, m=LifeSpan s1, SS=SCMPDS-Stop, s3=(Computation s1).m, s4=(Computation s2).m; A3: IJ=I ';' (G ';' J) by SCMPDS_4:46; A4: s1 is halting by A2,Def3; A5: IC s3 in dom pI by A1,Def2; consider n such that A6: inspos n= IC s3 by SCMPDS_3:32; card pI=card I + 1 by SCMPDS_5:7; then n < card I + 1 by A5,A6,SCMPDS_4:1; then A7: n <= card I by INT_1:20; A8: dom pI c= dom pJ by SCMPDS_5:16; set JS=G ';' J ';' SS; A9: pJ =I ';' (G ';' J) ';' SS by SCMPDS_4:def 7 .=I ';' JS by SCMPDS_4:46; A10: IsI c= s1 by FUNCT_4:26; pI c= IsI by SCMPDS_4:9; then A11: pI c= s1 by A10,XBOOLE_1:1; I c= I ';' SS by SCMPDS_4:40; then I c= pI by SCMPDS_4:def 7; then I c= s1 by A11,XBOOLE_1:1; then A12: I c= s3 by AMI_3:38; A13: IsJ c= s2 by FUNCT_4:26; pJ c= IsJ by SCMPDS_4:9; then A14: pJ c= s2 by A13,XBOOLE_1:1; I c= pJ by A9,SCMPDS_4:40; then I c= s2 by A14,XBOOLE_1:1; then A15: I c= s4 by AMI_3:38; per cases; suppose IC s3 <> inspos card I; then n < card I by A6,A7,REAL_1:def 5; then A16: IC s3 in dom I by A6,SCMPDS_4:1; A17: halt SCMPDS=CurInstr s3 by A4,SCM_1:def 2 .=s3.IC s3 by AMI_1:def 17 .=I.IC s3 by A12,A16,GRFUNC_1:8 .=s4.IC s3 by A15,A16,GRFUNC_1:8 .=s4.IC s4 by A1,A2,Th39 .=CurInstr s4 by AMI_1:def 17; then A18: s2 is halting by AMI_1:def 20; hence IJ is_halting_on s by A3,Def3; now let k be Nat; set C1k=IC (Computation s1).k, C2k=IC (Computation s2).k; per cases; suppose A19:k <= m; C1k in dom pI by A1,Def2; then C1k in dom pJ by A8; hence C2k in dom pJ by A1,A2,A19,Th39; suppose A20:k > m; set m2=LifeSpan s2; A21: m2 <= m by A17,A18,SCM_1:def 2; then k >= m2 by A20,AXIOMS:22; then C2k=IC (Computation s2).m2 by A18,Th3 .=IC (Computation s1).m2 by A1,A2,A21,Th39; then C2k in dom pI by A1,Def2; hence C2k in dom pJ by A8; end; hence IJ is_closed_on s by A3,Def2; suppose IC s3 =inspos card I; then A22: IC s4=inspos card I by A1,A2,Th39; A23: card (G ';' J) = card G + card J by SCMPDS_4:45 .=1 + card J by Th32; A24: card pJ = card I + card JS by A9,SCMPDS_4:45; A25: JS =G ';' (J ';' SS) by SCMPDS_4:46; then A26: card JS =card G + card (J ';' SS) by SCMPDS_4:45 .= 1 + card (J ';' SS) by Th32 .= card J + 1 + 1 by SCMPDS_4:45,74; then 0 + 1 <= card JS by NAT_1:29; then A27: 0 < card JS by NAT_1:38; then card I + 0 < card pJ by A24,REAL_1:53; then A28: inspos card I in dom pJ by SCMPDS_4:1; A29: inspos 0 in dom JS by A27,SCMPDS_4:1; A30: inspos 0 in dom G by Th33; A31: CurInstr s4= s4.inspos card I by A22,AMI_1:def 17 .= s2.inspos card I by AMI_1:54 .= (I ';' JS).inspos(0+card I) by A9,A14,A28,GRFUNC_1:8 .= (I ';' JS).(inspos 0 + card I) by SCMPDS_3:def 3 .= JS.inspos 0 by A29,SCMPDS_4:38 .=G.inspos 0 by A25,A30,SCMPDS_4:37 .=goto (card J + 1) by Th33; card J + 1 < card JS by A26,NAT_1:38; then A32: inspos (card J + 1) in dom JS by SCMPDS_4:1; card pJ = card I + (card J + (1 + 1)) by A24,A26,XCMPLX_1:1 .= card I + card J + (1 + 1) by XCMPLX_1:1 .= card I + card J + 1 + 1 by XCMPLX_1:1; then A33: card I + card J + 1 < card pJ by NAT_1:38; then A34: inspos (card I + card J + 1) in dom pJ by SCMPDS_4:1; A35: JS.inspos (card J + 1)=JS.inspos(0+card (G ';' J)) by A23 .=JS.(inspos 0+ card (G ';' J)) by SCMPDS_3:def 3 .=halt SCMPDS by SCMPDS_4:38,73; A36: IC (Computation s2).(m + 1) = IC Following s4 by AMI_1:def 19 .= IC Exec (goto (card J + 1),s4) by A31,AMI_1:def 18 .= Exec (goto (card J + 1),s4).IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card J +1) by SCMPDS_2:66 .=inspos (card I + (card J + 1)) by A22,Th23 .=inspos (card I + card J + 1) by XCMPLX_1:1; then A37: CurInstr (Computation s2).(m + 1) = (Computation s2).(m+ 1).inspos (card I + card J + 1) by AMI_1:def 17 .= s2.inspos (card I + card J + 1) by AMI_1:54 .= (I ';' JS).inspos (card I + card J + 1) by A9,A14,A34,GRFUNC_1:8 .= (I ';' JS).inspos(card I+(card J+1)) by XCMPLX_1:1 .= (I ';' JS).(inspos (card J+1)+card I) by SCMPDS_3:def 3 .= halt SCMPDS by A32,A35,SCMPDS_4:38; then A38: s2 is halting by AMI_1:def 20; hence IJ is_halting_on s by A3,Def3; now let k be Nat; set C1k=IC (Computation s1).k, C2k=IC (Computation s2).k; per cases; suppose A39:k <= m; C1k in dom pI by A1,Def2; then C1k in dom pJ by A8; hence C2k in dom pJ by A1,A2,A39,Th39; suppose k > m; then A40: k >= m+1 by INT_1:20; set m2=LifeSpan s2; A41: m2 <= m+1 by A37,A38,SCM_1:def 2; then k >= m2 by A40,AXIOMS:22; then C2k=IC (Computation s2).m2 by A38,Th3 .=inspos (card I + card J + 1) by A36,A38,A41,Th3; hence C2k in dom pJ by A33,SCMPDS_4:1; end; hence IJ is_closed_on s by A3,Def2; end; theorem Th45: :: SP4_88,Th27 for I being shiftable Program-block st Initialized stop I c= s1 & I is_closed_on s1 for n being Nat st Shift(stop I,n) c= s2 & IC s2 = inspos n & s1 | SCM-Data-Loc = s2 | SCM-Data-Loc for i being Nat holds IC (Computation s1).i + n = IC (Computation s2).i & CurInstr ((Computation s1).i) = CurInstr ((Computation s2).i) & (Computation s1).i | SCM-Data-Loc = (Computation s2).i | SCM-Data-Loc proof let I be shiftable Program-block; set SI=stop I, II = Initialized SI; assume A1: II c= s1 & I is_closed_on s1; let n be Nat; assume that A2: Shift(SI,n) c= s2 and A3: IC s2 = inspos n and A4: s1 | SCM-Data-Loc = s2 | SCM-Data-Loc; A5: s1=s1 +* II by A1,AMI_5:10; set C1 = Computation s1; set C2 = Computation s2; let i be Nat; defpred P[Nat] means IC C1.$1 + n = IC C2.$1 & CurInstr (C1.$1) = CurInstr (C2.$1) & C1.$1 | SCM-Data-Loc = C2.$1 | SCM-Data-Loc; A6: II= SI +* Start-At inspos 0 by SCMPDS_4:def 2; dom SI misses dom Start-At inspos 0 by SCMPDS_4:54; then A7: SI c= II by A6,FUNCT_4:33; then A8: dom SI c= dom II by GRFUNC_1:8; A9: inspos 0 in dom SI by SCMPDS_4:75; A10: P[0] proof A11: IC SCMPDS in dom II by SCMPDS_4:7; inspos 0 + n in dom Shift(SI,n) by A9,SCMPDS_4:76; then A12: inspos (0 + n) in dom Shift(SI,n) by SCMPDS_3:def 3; IC C1.0 = IC s1 by AMI_1:def 19 .= s1.IC SCMPDS by AMI_1:def 15 .= II.IC SCMPDS by A1,A11,GRFUNC_1:8 .= inspos 0 by SCMPDS_4:29; hence IC C1.0 + n = inspos (0 + n) by SCMPDS_3:def 3 .= IC C2.0 by A3,AMI_1:def 19; A13: s1.IC s1 = s1.(s1.IC SCMPDS) by AMI_1:def 15 .= s1.((II).IC SCMPDS) by A1,A11,GRFUNC_1:8 .= s1.inspos 0 by SCMPDS_4:29 .= II.inspos 0 by A1,A8,A9,GRFUNC_1:8 .= SI.inspos 0 by A7,A9,GRFUNC_1:8; thus CurInstr (C1.0) = CurInstr s1 by AMI_1:def 19 .= s1.IC s1 by AMI_1:def 17 .= Shift(SI,n).(inspos 0 + n) by A9,A13,SCMPDS_3:37 .= Shift(SI,n).inspos (0 + n) by SCMPDS_3:def 3 .= s2.IC s2 by A2,A3,A12,GRFUNC_1:8 .= CurInstr s2 by AMI_1:def 17 .= CurInstr (C2.0) by AMI_1:def 19; thus C1.0 | SCM-Data-Loc = s2 | SCM-Data-Loc by A4,AMI_1:def 19 .= C2.0 | SCM-Data-Loc by AMI_1:def 19; end; A14: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume A15: P[k]; set i = CurInstr C1.k; A16: C1.(k + 1) = Following C1.k by AMI_1:def 19 .= Exec(CurInstr C1.k,C1.k) by AMI_1:def 18; A17: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18; A18: IC C1.k in dom SI by A1,A5,Def2; A19: i = C1.k.(IC C1.k) by AMI_1:def 17 .= s1.IC C1.k by AMI_1:54 .= II.IC C1.k by A1,A8,A18,GRFUNC_1:8 .= SI.IC C1.k by A7,A18,GRFUNC_1:8; consider m such that A20: IC C1.k =inspos m by SCMPDS_3:32; A21: InsCode i <> 1 & InsCode i <> 3 & i valid_at m by A18,A19,A20,SCMPDS_4:def 12; hence A22: IC C1.(k + 1) + n = IC C2.(k + 1) by A15,A16,A17,A20,SCMPDS_4:83; set l = IC C1.(k + 1); A23: IC C1.(k + 1) in dom SI by A1,A5,Def2; A24: CurInstr C1.(k + 1) = C1.(k + 1).l by AMI_1:def 17 .= s1.l by AMI_1:54 .= II.l by A1,A8,A23,GRFUNC_1:8 .= SI.l by A7,A23,GRFUNC_1:8; A25: IC C2.(k + 1) in dom Shift(SI,n) by A22,A23,SCMPDS_4:76; thus CurInstr C1.(k + 1) = Shift(SI,n).(IC C2.(k + 1)) by A22,A23,A24,SCMPDS_3:37 .= s2.IC C2.(k + 1) by A2,A25,GRFUNC_1:8 .= C2.(k + 1).IC C2.(k + 1) by AMI_1:54 .= CurInstr C2.(k + 1) by AMI_1:def 17; thus C1.(k + 1) | SCM-Data-Loc = C2.(k + 1) | SCM-Data-Loc by A15,A16,A17,A20,A21,SCMPDS_4:83; end; for k being Nat holds P[k] from Ind(A10,A14); hence thesis; end; Lm3: ::SHIFT for s being State of SCMPDS, I,J being shiftable Program-block,n being Nat holds I ';' Goto n ';' J is shiftable proof let s be State of SCMPDS, I,J be shiftable Program-block,n be Nat; I ';' Goto n ';' J = I ';' Load (goto n) ';' J by Def1; hence thesis; end; theorem Th46: ::SCMFSA8A:61 for s being State of SCMPDS,I being No-StopCode Program-block, J being Program-block st I is_closed_on s & I is_halting_on s holds IC IExec(I ';' Goto (card J + 1) ';' J,s) =inspos (card I + card J + 1) proof let s be State of SCMPDS,I be No-StopCode Program-block, J be Program-block; set m= LifeSpan(s +* Initialized stop I)+1, G=Goto (card J + 1), s2 = s +* Initialized stop (I ';' G ';' J); assume A1: I is_closed_on s & I is_halting_on s; then s2 is halting & LifeSpan s2 = m by Lm2; then IC Result s2 = IC (Computation s2).m by SCMFSA6B:16 .= inspos (card I + card J + 1) by A1,Lm2; hence IC IExec(I ';' G ';' J, s) = inspos (card I + card J + 1) by SCMPDS_5:22; end; theorem Th47: ::SCMFSA8A:62 for s being State of SCMPDS,I being No-StopCode Program-block, J being Program-block st I is_closed_on s & I is_halting_on s holds IExec(I ';' Goto (card J + 1) ';' J,s) = IExec(I,s) +* Start-At inspos (card I + card J + 1) proof let s be State of SCMPDS,I be No-StopCode Program-block, J be Program-block; set s1= s +* Initialized stop I, m= LifeSpan s1+1, G=Goto (card J + 1), s2 = s +* Initialized stop (I ';' G ';' J), l= inspos (card I + card J + 1); assume A1: I is_closed_on s & I is_halting_on s; A2: dom (s | A) = A by Th1; A3: s1 is halting by A1,Def3; s2 is halting & LifeSpan s2 = m by A1,Lm2; then A4: Result s2 = (Computation s2).m by SCMFSA6B:16; then (Result s2) | D = (Computation s1).(LifeSpan s1) | D by A1,Lm2; then A5: (Result s2) | D = (Result s1) | D by A3,SCMFSA6B:16 .= (Result s1 +* Start-At l) | D by Th7; IC Result s2 = l by A1,A4,Lm2 .= IC (Result s1 +* Start-At l) by AMI_5:79; then Result s2,Result s1 +* Start-At l equal_outside A by A5,Th4; then A6: Result s2 +* s | A = Result s1 +* Start-At l +* s | A by A2,SCMFSA6A: 13; A7: dom (s | A) misses dom Start-At l by Th10; thus IExec(I ';' G ';' J,s) = Result s2 +* s | A by SCMPDS_4:def 8 .= Result s1 +* (Start-At l +* s | A) by A6,FUNCT_4:15 .= Result s1 +* (s | A +* Start-At l) by A7,FUNCT_4:36 .= Result s1 +* s | A +* Start-At l by FUNCT_4:15 .= IExec(I,s) +* Start-At l by SCMPDS_4:def 8; end; theorem Th48: for s being State of SCMPDS,I being No-StopCode Program-block st I is_closed_on s & I is_halting_on s holds IC IExec(I,s) = inspos card I proof let s be State of SCMPDS,I be No-StopCode Program-block; set s1=s+*Initialized stop(I); assume A1: I is_closed_on s & I is_halting_on s; then A2: s1 is halting by Def3; thus IC IExec(I,s) = IC Result s1 by SCMPDS_5:22 .= IC (Computation s1).(LifeSpan s1) by A2,SCMFSA6B:16 .=inspos card I by A1,Th43; end; begin :: The construction of conditional statements definition let a be Int_position,k be Integer; let I,J be Program-block; func if=0(a,k,I,J) -> Program-block equals :Def4: (a,k)<>0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J; coherence; func if>0(a,k,I,J) -> Program-block equals :Def5: (a,k)<=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J; coherence; func if<0(a,k,I,J) -> Program-block equals :Def6: (a,k)>=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J; coherence; end; definition let a be Int_position,k be Integer; let I be Program-block; func if=0(a,k,I) -> Program-block equals :Def7: (a,k)<>0_goto (card I +1) ';' I; coherence; func if<>0(a,k,I) -> Program-block equals :Def8: (a,k)<>0_goto 2 ';' goto (card I+1) ';' I; coherence; func if>0(a,k,I) -> Program-block equals :Def9: (a,k)<=0_goto (card I +1) ';' I; coherence; func if<=0(a,k,I) -> Program-block equals :Def10: (a,k)<=0_goto 2 ';' goto (card I+1) ';' I; coherence; func if<0(a,k,I) -> Program-block equals :Def11: (a,k)>=0_goto (card I +1) ';' I; coherence; func if>=0(a,k,I) -> Program-block equals :Def12: (a,k)>=0_goto 2 ';' goto (card I+1) ';' I; coherence; end; Lm4: card (i ';' I ';' Goto n ';' J) = card I + card J +2 proof set G=Goto n; thus card (i ';' I ';' G ';' J) =card (i ';' I ';' G) + card J by SCMPDS_4:45 .=card (i ';' I) + card G + card J by SCMPDS_4:45 .=card (i ';' I) + 1 + card J by Th32 .=card I +1 +1 +card J by Th15 .=card I +(1 +1)+card J by XCMPLX_1:1 .=card I + card J +2 by XCMPLX_1:1; end; begin :: The computation of "if var=0 then block1 else block2" theorem Th49: ::S8B_14 card if=0(a,k1,I,J) = card I + card J + 2 proof if=0(a,k1,I,J)= (a,k1)<>0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def4; hence thesis by Lm4; end; theorem ::LmT5 inspos 0 in dom if=0(a,k1,I,J) & inspos 1 in dom if=0(a,k1,I,J) proof set ci=card if=0(a,k1,I,J); ci=card I + card J +2 by Th49; then 2 <= ci by NAT_1:37; then 0 < ci & 1 < ci by AXIOMS:22; hence thesis by SCMPDS_4:1; end; Lm5: (i ';' I ';' J ';' K).inspos 0=i proof A1: i ';' I ';' J ';' K =i ';' (I ';' J) ';' K by SCMPDS_4:50 .=i ';' (I ';' J ';' K) by SCMPDS_4:50 .=Load i ';' (I ';' J ';' K) by SCMPDS_4:def 4; inspos 0 in dom Load i by SCMPDS_5:2; hence (i ';' I ';' J ';' K).inspos 0 =(Load i).inspos 0 by A1,SCMPDS_4:37 .=i by SCMPDS_5:4; end; theorem ::Lm6 if=0(a,k1,I,J).inspos 0 = (a,k1)<>0_goto (card I + 2) proof if=0(a,k1,I,J)= (a,k1)<>0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def4; hence thesis by Lm5; end; Lm6: Shift(stop I,1) c= (Computation (s+* Initialized stop(i ';' I))).n proof set pI=stop I, iI=i ';' I, piI=stop iI, IsiI=Initialized piI, s3=s+*IsiI; A1: IsiI c= s3 by FUNCT_4:26; A2: card Load i=1 by SCMPDS_5:6; iI=(Load i) ';' I by SCMPDS_4:def 4; then A3: Shift(pI,1) c= piI by A2,Th24; piI c= s3 by A1,SCMPDS_4:57; then Shift(pI,1) c= s3 by A3,XBOOLE_1:1; hence Shift(pI,1) c= (Computation s3).n by AMI_3:38; end; Lm7: Shift(stop I,2) c= (Computation (s+* Initialized stop(i ';' j ';' I))).n proof set pI=stop I, pjI=stop (i ';' j ';' I), IsjI=Initialized pjI, s3=s+*IsjI; A1: IsjI c= s3 by FUNCT_4:26; card (i ';' j)=card (Load i ';' Load j) by SCMPDS_4:def 6 .=card Load i + card Load j by SCMPDS_4:45 .=1+ card Load j by SCMPDS_5:6 .=1+1 by SCMPDS_5:6; then A2: Shift(pI,2) c= pjI by Th24; pjI c= s3 by A1,SCMPDS_4:57; then Shift(pI,2) c= s3 by A2,XBOOLE_1:1; hence Shift(pI,2) c= (Computation s3).n by AMI_3:38; end; theorem Th52: ::S8B_18 for s being State of SCMPDS, I,J being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)= 0 & I is_closed_on s & I is_halting_on s holds if=0(a,k1,I,J) is_closed_on s & if=0(a,k1,I,J) is_halting_on s proof let s be State of SCMPDS, I,J be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b = 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set G=Goto (card J+1); set I2 = I ';' G ';' J, IF=if=0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, pI2=stop I2, II2= Initialized pI2, s2 = s +* II2, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<>0_goto (card I + 2); A4: II2 c= s2 by FUNCT_4:26; I2 is_halting_on s by A2,A3,Th44; then A5: s2 is halting by Def3; A6: I2 is_closed_on s by A2,A3,Th44; then A7: I2 is_closed_on s2 by Th38; A8: inspos 0 in dom pIF by SCMPDS_4:75; A9: IF = i ';' I ';' G ';' J by Def4 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' I2 by SCMPDS_4:50; A10: IC s3 =inspos 0 by Th21; A11: CurInstr s3 = i by A9,Th22; A12: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A11,AMI_1:def 18; A13: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A14: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A15: s3.DataLoc(s3.a,k1)=s3.b by A13,FUNCT_4:12 .=0 by A1,A14,FUNCT_4:12; A16: card pIF = card IF +1 by SCMPDS_5:7 .= card I2 +1+1 by A9,Th15; A17: Shift(pI2,1) c= s4 by A9,Lm6; A18: I2 is shiftable by Lm3; A19: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A12,A15,SCMPDS_2:67 .= inspos(0+1) by A10,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A20: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A20,SCMPDS_4:23 .= s4.a by A12,SCMPDS_2:67; end; then A21: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A7,A17,A18,A19,A21,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A22: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A23: k1 + 1 = k by NAT_1:22; consider m such that A24: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A25: card pIF = card pI2+1 by A16,SCMPDS_5:7; inspos m in dom pI2 by A6,A24,Def2; then m < card pI2 by SCMPDS_4:1; then A26: m+1 < card pIF by A25,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A23,AMI_1:51 .= IC (Computation s2).k1 + 1 by A4,A7,A17,A18,A19,A21,Th45 .= inspos (m + 1) by A24,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A26,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A8,A10,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A22,Def3; end; theorem Th53: ::S8B_16 for s being State of SCMPDS,I being Program-block,J being shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)<> 0 & J is_closed_on s & J is_halting_on s holds if=0(a,k1,I,J) is_closed_on s & if=0(a,k1,I,J) is_halting_on s proof let s be State of SCMPDS,I be Program-block,J be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <> 0; assume A2: J is_closed_on s; assume A3: J is_halting_on s; set pJ=stop J, IsJ=Initialized pJ, s1 = s +* IsJ, IF=if=0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<>0_goto (card I + 2); set G =Goto (card J+1), iG=i ';' I ';' G; A4: IsJ c= s1 by FUNCT_4:26; A5: s1 is halting by A3,Def3; A6: J is_closed_on s1 by A2,Th38; A7: IF = iG ';' J by Def4; A8: IF = i ';' I ';' G ';' J by Def4 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' (I ';' G ';' J) by SCMPDS_4:50; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: card iG = card (i ';' I) + card G by SCMPDS_4:45 .=card (i ';' I) + 1 by Th32 .=card I +1 +1 by Th15 .=card I +(1 +1) by XCMPLX_1:1; A13: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A14: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A15: s3.DataLoc(s3.a,k1)=s3.b by A13,FUNCT_4:12 .= s.b by A14,FUNCT_4:12; A16: IsIF c= s3 by FUNCT_4:26; s1,s3 equal_outside A by SCMPDS_4:36; then A17: s1 | D = s3 | D by SCMPDS_4:24; now let a; thus s1.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:67; end; then A18: s1 | D = s4 | D by SCMPDS_4:23; A19: Shift(pJ,card I+2) c= pIF by A7,A12,Th24; pIF c= s3 by A16,SCMPDS_4:57; then Shift(pJ,card I+2) c= s3 by A19,XBOOLE_1:1; then A20: Shift(pJ,card I+2) c= s4 by AMI_3:38; A21: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 2) by A1,A11,A15,SCMPDS_2:67 .= inspos(0+(card I + 2)) by A9,Th23; CurInstr (Computation s3).(LifeSpan s1 + 1) =CurInstr (Computation s4).LifeSpan s1 by AMI_1:51 .=CurInstr (Computation s1).LifeSpan s1 by A4,A6,A18,A20,A21,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A22: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A23: k1 + 1 = k by NAT_1:22; consider m such that A24: inspos m = IC (Computation s1).k1 by SCMPDS_3:32; A25: card pJ = card J + 1 by SCMPDS_5:7; A26: card pIF = card IF+1 by SCMPDS_5:7 .=card I +2 +card J +1 by A7,A12,SCMPDS_4:45 .=card I +2 + card pJ by A25,XCMPLX_1:1; inspos m in dom pJ by A2,A24,Def2; then m < card pJ by SCMPDS_4:1; then A27: m + (card I + 2) < card pJ + (card I + 2) by REAL_1:53; IC C3.k = IC (Computation s4).k1 by A23,AMI_1:51 .= IC (Computation s1).k1 + (card I + 2) by A4,A6,A18,A20,A21,Th45 .= inspos (m + (card I + 2)) by A24,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A26,A27,SCMPDS_4:1; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A9,SCMPDS_4:75; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A22,Def3; end; theorem Th54: ::E,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, J being shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)= 0 & I is_closed_on s & I is_halting_on s holds IExec(if=0(a,k1,I,J),s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, J be shiftable Program-block,a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b = 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set G=Goto (card J+1); set I2 = I ';' G ';' J, IF=if=0(a,k1,I,J), IsIF=Initialized stop IF, pI2=stop I2, II2= Initialized pI2, s2 = s +* II2, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<>0_goto (card I + 2); set SAl= Start-At inspos (card I + card J + 2); A4: II2 c= s2 by FUNCT_4:26; I2 is_halting_on s by A2,A3,Th44; then A5: s2 is halting by Def3; I2 is_closed_on s by A2,A3,Th44; then A6: I2 is_closed_on s2 by Th38; A7: IF = i ';' I ';' G ';' J by Def4 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' I2 by SCMPDS_4:50; A8: IC s3 =inspos 0 by Th21; A9: CurInstr s3 = i by A7,Th22; A10: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A9,AMI_1:def 18; A11: dom (s | A) = A by Th1; A12: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A14: s3.DataLoc(s3.a,k1)=s3.b by A12,FUNCT_4:12 .=0 by A1,A13,FUNCT_4:12; A15: Shift(pI2,1) c= s4 by A7,Lm6; A16: I2 is shiftable by Lm3; A17: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A10,A14,SCMPDS_2:67 .= inspos(0+1) by A8,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A18: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A18,SCMPDS_4:23 .= s4.a by A10,SCMPDS_2:67; end; then A19: s2 | D = s4 | D by SCMPDS_4:23; A20: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A17,A19,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A21: s3 is halting by AMI_1:def 20; now let l be Nat; assume A22: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A9,Th29; suppose l <> 0; then consider n such that A23: l = n + 1 by NAT_1:22; A24: n < LifeSpan s2 by A22,A23,AXIOMS:24; assume A25: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A15,A16,A17,A19,Th45 .= halt SCMPDS by A23,A25,AMI_1:51; hence contradiction by A5,A24,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A26: LifeSpan s3 = LifeSpan s2 + 1 by A20,A21,SCM_1:def 2; A27: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A15,A16,A17,A19,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A21,A26,SCMFSA6B:16; A28: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I2,s) +* SAl) by AMI_3:36; now let x be set; A29: IExec(I2,s) = Result s2 +* s | A by SCMPDS_4:def 8; A30: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A31: x in dom IExec(IF,s); A32: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A31,SCMPDS_4:20; suppose A33: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A34: not x in dom SAl by A32,TARSKI:def 1; A35: not x in dom (s | A) by A11,A33,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A30,FUNCT_4:12 .= (Result s2).x by A27,A33,SCMPDS_4:23 .= IExec(I2,s).x by A29,A35,FUNCT_4:12 .= (IExec(I2,s) +* SAl).x by A34,FUNCT_4:12; suppose A36: x = IC SCMPDS; then A37: x in dom SAl by A32,TARSKI:def 1; A38: not x in dom (s | A) by A11,A36,AMI_1:48; A39: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I2,s).IC SCMPDS by A29,A36,A38,FUNCT_4:12 .= IC IExec(I2,s) by AMI_1:def 15 .= inspos (card I + card J + 1) by A2,A3,Th46; thus IExec(IF,s).x = (Result s3).x by A30,A38,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A21,A26,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A36,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 1 by A4,A6,A15,A16,A17,A19,Th45 .= IC Result s2 + 1 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I + card J + 1) + 1)).IC SCMPDS by A39,AMI_3:50 .= (Start-At inspos (card I + card J + 1 + 1)).IC SCMPDS by SCMPDS_3:def 3 .= (Start-At inspos (card I + card J + (1 + 1))).IC SCMPDS by XCMPLX_1:1 .= (IExec(I2,s) +* SAl).x by A36,A37,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I2,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I2,s) +* SAl by A28,FUNCT_1:9 .= IExec(I,s) +* Start-At inspos (card I + card J + 1) +* Start-At inspos (card I + card J + 2) by A2,A3,Th47 .= IExec(I,s) +* SAl by Th14; end; theorem Th55: ::E,SCM8B_17 for s being State of SCMPDS,I being Program-block,J being No-StopCode shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)<> 0 & J is_closed_on s & J is_halting_on s holds IExec(if=0(a,k1,I,J),s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) proof let s be State of SCMPDS,I be Program-block,J be No-StopCode shiftable Program-block,a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <> 0; assume A2: J is_closed_on s; assume A3: J is_halting_on s; set pJ=stop J, IsJ=Initialized pJ, s1 = s +* IsJ, IF=if=0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<>0_goto (card I + 2); set G =Goto (card J+1), iG=i ';' I ';' G; set SAl=Start-At inspos (card I + card J + 2); A4: IsJ c= s1 by FUNCT_4:26; A5: s1 is halting by A3,Def3; A6: J is_closed_on s1 by A2,Th38; A7: IF = iG ';' J by Def4; A8: IF = i ';' I ';' G ';' J by Def4 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' (I ';' G ';' J) by SCMPDS_4:50; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: card iG = card (i ';' I) + card G by SCMPDS_4:45 .=card (i ';' I) + 1 by Th32 .=card I +1 +1 by Th15 .=card I +(1 +1) by XCMPLX_1:1; A13: dom (s | A) = A by Th1; A14: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A15: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A16: s3.DataLoc(s3.a,k1)=s3.b by A14,FUNCT_4:12 .= s.b by A15,FUNCT_4:12; A17: IsIF c= s3 by FUNCT_4:26; A18: Shift(pJ,card I+2) c= pIF by A7,A12,Th24; pIF c= s3 by A17,SCMPDS_4:57; then Shift(pJ,card I+2) c= s3 by A18,XBOOLE_1:1; then A19: Shift(pJ,card I+2) c= s4 by AMI_3:38; A20: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 2) by A1,A11,A16,SCMPDS_2:67 .= inspos(0+(card I + 2)) by A9,Th23; s1,s3 equal_outside A by SCMPDS_4:36; then A21: s1 | D = s3 | D by SCMPDS_4:24; now let a; thus s1.a = s3.a by A21,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:67; end; then A22: s1 | D = s4 | D by SCMPDS_4:23; A23: CurInstr (Computation s3).(LifeSpan s1 + 1) =CurInstr (Computation s4).LifeSpan s1 by AMI_1:51 .=CurInstr (Computation s1).LifeSpan s1 by A4,A6,A19,A20,A22,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A24: s3 is halting by AMI_1:def 20; now let l be Nat; assume A25: l < LifeSpan s1 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A10,Th29; suppose l <> 0; then consider n such that A26: l = n + 1 by NAT_1:22; A27: n < LifeSpan s1 by A25,A26,AXIOMS:24; assume A28: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s1).n = CurInstr (Computation s4).n by A4,A6,A19,A20,A22,Th45 .= halt SCMPDS by A26,A28,AMI_1:51; hence contradiction by A5,A27,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s1 + 1 <= l; then A29: LifeSpan s3 = LifeSpan s1 + 1 by A23,A24,SCM_1:def 2; A30: (Result s1) | D = (Computation s1).(LifeSpan s1) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s1) | D by A4,A6,A19,A20,A22,Th45 .= (Computation s3).(LifeSpan s1 + 1) | D by AMI_1:51 .= (Result s3) | D by A24,A29,SCMFSA6B:16; A31: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(J,s) +* SAl) by AMI_3:36; now let x be set; A32: IExec(J,s) = Result s1 +* s | A by SCMPDS_4:def 8; A33: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A34: x in dom IExec(IF,s); A35: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A34,SCMPDS_4:20; suppose A36: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A37: not x in dom SAl by A35,TARSKI:def 1; A38: not x in dom (s | A) by A13,A36,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A33,FUNCT_4:12 .= (Result s1).x by A30,A36,SCMPDS_4:23 .= IExec(J,s).x by A32,A38,FUNCT_4:12 .= (IExec(J,s) +* SAl).x by A37,FUNCT_4:12; suppose A39: x = IC SCMPDS; then A40: x in dom SAl by A35,TARSKI:def 1; A41: not x in dom (s | A) by A13,A39,AMI_1:48; A42: IC Result s1 = (Result s1).IC SCMPDS by AMI_1:def 15 .= IExec(J,s).IC SCMPDS by A32,A39,A41,FUNCT_4:12 .= IC IExec(J,s) by AMI_1:def 15 .= inspos (card J) by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A33,A41,FUNCT_4:12 .= (Computation s3).(LifeSpan s1 + 1).x by A24,A29,SCMFSA6B:16 .= (Computation s4).(LifeSpan s1).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s1) by A39,AMI_1:def 15 .= IC (Computation s1).(LifeSpan s1) + (card I + 2) by A4,A6,A19,A20,A22,Th45 .= IC Result s1 + (card I + 2) by A5,SCMFSA6B:16 .= (Start-At (inspos card J + (card I + 2))).IC SCMPDS by A42,AMI_3:50 .= (Start-At inspos (card I + 2+ card J)).IC SCMPDS by SCMPDS_3:def 3 .= SAl.IC SCMPDS by XCMPLX_1:1 .= (IExec(J,s) +* SAl).x by A39,A40,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(J,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(J,s) +* SAl by A31,FUNCT_1:9; end; definition let I,J be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if=0(a,k1,I,J) -> shiftable parahalting; correctness proof set IF=if=0(a,k1,I,J), IsIF=Initialized stop IF; set i = (a,k1)<>0_goto (card I + 2), G =Goto (card J+1); reconsider IJ=I ';' G ';' J as shiftable Program-block by Lm3; IF = i ';' I ';' G ';' J by Def4 .=i ';' (I ';' G) ';' J by SCMPDS_4:50 .=i ';' IJ by SCMPDS_4:50 .=Load i ';' IJ by SCMPDS_4:def 4; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; A3: J is_closed_on s & J is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1)= 0; then IF is_halting_on s by A2,Th52; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) <> 0; then IF is_halting_on s by A3,Th53; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I,J be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if=0(a,k1,I,J) -> No-StopCode; coherence proof if=0(a,k1,I,J) = (a,k1)<>0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def4; hence thesis; end; end; theorem ::E,S8B_21A for s being State of SCMPDS,I,J being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if=0(a,k1,I,J),s) = inspos (card I + card J + 2) proof let s be State of SCMPDS,I,J be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if=0(a,k1,I,J); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; A2: J is_closed_on s & J is_halting_on s by Th34,Th35; hereby per cases; suppose s.DataLoc(s.a,k1) = 0; then IExec(IF,s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) by A1,Th54; hence IC IExec(IF,s) = inspos (card I + card J + 2) by AMI_5:79; suppose s.DataLoc(s.a,k1) <> 0; then IExec(IF,s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) by A2,Th55; hence IC IExec(IF,s) = inspos (card I + card J + 2) by AMI_5:79; end; end; theorem ::E,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,J being shiftable Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1)= 0 holds IExec(if=0(a,k1,I,J),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,J be shiftable Program-block,a,b be Int_position, k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1)=0; then A2: IExec(if=0(a,k1,I,J),s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) by A1,Th54; not b in dom Start-At inspos (card I + card J + 2) by SCMPDS_4:59; hence IExec(if=0(a,k1,I,J),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::E,S8B_21C for s being State of SCMPDS,I being Program-block,J being No-StopCode parahalting shiftable Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1)<> 0 holds IExec(if=0(a,k1,I,J),s).b = IExec(J,s).b proof let s be State of SCMPDS,I be Program-block,J be No-StopCode parahalting shiftable Program-block,a,b be Int_position, k1 be Integer; A1: J is_closed_on s & J is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1)<>0; then A2: IExec(if=0(a,k1,I,J),s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) by A1,Th55; not b in dom Start-At inspos (card I + card J + 2) by SCMPDS_4:59; hence IExec(if=0(a,k1,I,J),s).b = IExec(J,s).b by A2,FUNCT_4:12; end; begin :: The computation of "if var=0 then block" theorem Th59: ::E,S8B_14 card if=0(a,k1,I) = card I + 1 proof thus card if=0(a,k1,I)=card ((a,k1)<>0_goto (card I +1) ';' I) by Def7 .=card I+1 by Th15; end; theorem inspos 0 in dom if=0(a,k1,I) proof set ci=card if=0(a,k1,I); ci=card I + 1 by Th59; then 0 < ci by NAT_1:19; hence thesis by SCMPDS_4:1; end; theorem ::Lm6 if=0(a,k1,I).inspos 0 = (a,k1)<>0_goto (card I + 1) proof if=0(a,k1,I)=(a,k1)<>0_goto (card I +1) ';' I by Def7; hence thesis by Th16; end; theorem Th62: ::E,S8B_18 for s being State of SCMPDS, I being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)= 0 & I is_closed_on s & I is_halting_on s holds if=0(a,k1,I) is_closed_on s & if=0(a,k1,I) is_halting_on s proof let s be State of SCMPDS, I be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b = 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<>0_goto (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: inspos 0 in dom pIF by SCMPDS_4:75; A8: IF = i ';' I by Def7; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A14: s3.DataLoc(s3.a,k1)=s3.b by A12,FUNCT_4:12 .=0 by A1,A13,FUNCT_4:12; A15: card pIF = card IF +1 by SCMPDS_5:7 .= card I +1+1 by A8,Th15; A16: Shift(pI,1) c= s4 by A8,Lm6; A17: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A11,A14,SCMPDS_2:67 .= inspos(0+1) by A9,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A18: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A18,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:67; end; then A19: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A16,A17,A19,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A21: k1 + 1 = k by NAT_1:22; consider m such that A22: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A23: card pIF = card pI+1 by A15,SCMPDS_5:7; inspos m in dom pI by A2,A22,Def2; then m < card pI by SCMPDS_4:1; then A24: m+1 < card pIF by A23,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A21,AMI_1:51 .= IC (Computation s2).k1 + 1 by A4,A6,A16,A17,A19,Th45 .= inspos (m + 1) by A22,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A24,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A7,A9,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A20,Def3; end; theorem Th63: ::E,S8B_16 for s being State of SCMPDS,I being Program-block, a being Int_position, k1 being Integer st s.DataLoc(s.a,k1)<> 0 holds if=0(a,k1,I) is_closed_on s & if=0(a,k1,I) is_halting_on s proof let s be State of SCMPDS,I be Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <> 0; set IF=if=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<>0_goto (card I + 1); A2: IF = i ';' I by Def7; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A7: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A8: s3.DataLoc(s3.a,k1)=s3.b by A6,FUNCT_4:12 .= s.b by A7,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: card IF=card I+1 by Th59; then A11: inspos(card I+1) in dom pIF by Th25; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 1) by A1,A5,A8,SCMPDS_2:67 .= inspos(0+(card I + 1)) by A3,Th23; s4.inspos(card I+1) = pIF.inspos(card I+1) by A9,A11,GRFUNC_1:8 .=halt SCMPDS by A10,Th25; then A13: CurInstr s4 = halt SCMPDS by A12,AMI_1:def 17; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then 1+0 <= k by INT_1:20; hence IC C3.k in dom pIF by A11,A12,A13,AMI_1:52; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A3,SCMPDS_4:75; end; hence IF is_closed_on s by Def2; s3 is halting by A13,AMI_1:def 20; hence IF is_halting_on s by Def3; end; theorem Th64: ::E,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)= 0 & I is_closed_on s & I is_halting_on s holds IExec(if=0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 1) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b = 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if=0(a,k1,I), IsIF=Initialized stop IF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<>0_goto (card I + 1); set SAl=Start-At inspos (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: IF = i ';' I by Def7; A8: IC s3 =inspos 0 by Th21; A9: CurInstr s3 = i by A7,Th22; A10: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A9,AMI_1:def 18; A11: dom (s | A) = A by Th1; A12: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A14: s3.DataLoc(s3.a,k1)=s3.b by A12,FUNCT_4:12 .=0 by A1,A13,FUNCT_4:12; A15: Shift(pI,1) c= s4 by A7,Lm6; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A10,A14,SCMPDS_2:67 .= inspos(0+1) by A8,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A10,SCMPDS_2:67; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; A19: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume A21: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A9,Th29; suppose l <> 0; then consider n such that A22: l = n + 1 by NAT_1:22; A23: n < LifeSpan s2 by A21,A22,AXIOMS:24; assume A24: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A22,A24,AMI_1:51; hence contradiction by A5,A23,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A25: LifeSpan s3 = LifeSpan s2 + 1 by A19,A20,SCM_1:def 2; A26: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A15,A16,A18,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A20,A25,SCMFSA6B:16; A27: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I,s) +* SAl) by AMI_3:36; now let x be set; A28: IExec(I,s) = Result s2 +* s | A by SCMPDS_4:def 8; A29: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A30: x in dom IExec(IF,s); A31: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A30,SCMPDS_4:20; suppose A32: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A33: not x in dom SAl by A31,TARSKI:def 1; A34: not x in dom (s | A) by A11,A32,SCMPDS_2:53; x in dom Result s3 & not x in dom (s | A) by A11,A32,SCMPDS_2:49,53; hence IExec(IF,s).x = (Result s3).x by A29,FUNCT_4:12 .= (Result s2).x by A26,A32,SCMPDS_4:23 .= IExec(I,s).x by A28,A34,FUNCT_4:12 .= (IExec(I,s) +* SAl).x by A33,FUNCT_4:12; suppose A35: x = IC SCMPDS; then A36: x in dom SAl by A31,TARSKI:def 1; A37: not x in dom (s | A) by A11,A35,AMI_1:48; A38: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I,s).IC SCMPDS by A28,A35,A37,FUNCT_4:12 .= IC IExec(I,s) by AMI_1:def 15 .= inspos card I by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A29,A37,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A20,A25,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A35,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 1 by A4,A6,A15,A16,A18,Th45 .= IC Result s2 + 1 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I) + 1)).IC SCMPDS by A38,AMI_3:50 .= SAl.IC SCMPDS by SCMPDS_3:def 3 .= (IExec(I,s) +* SAl).x by A35,A36,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I,s) +* SAl by A27,FUNCT_1:9; end; Lm8: :: Th46 (s +* Start-At loc).IC SCMPDS = loc proof thus (s +* Start-At loc).IC SCMPDS=IC (s +* Start-At loc) by AMI_1:def 15 .=loc by AMI_5:79; end; theorem Th65: ::E,SCM8B_17 for s being State of SCMPDS,I being Program-block,a being Int_position, k1 being Integer st s.DataLoc(s.a,k1)<> 0 holds IExec(if=0(a,k1,I),s) = s +* Start-At inspos (card I + 1) proof let s be State of SCMPDS,I be Program-block,a be Int_position, k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <> 0; set IF=if=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<>0_goto (card I + 1); set SAl=Start-At inspos (card I + 1); A2: IF = i ';' I by Def7; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: dom (s | A) = A by Th1; A7: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A8: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A9: s3.DataLoc(s3.a,k1)=s3.b by A7,FUNCT_4:12 .= s.b by A8,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then pIF c= s3 by SCMPDS_4:57; then A10: pIF c= s4 by AMI_3:38; A11: card IF=card I+1 by Th59; then A12: inspos(card I+1) in dom pIF by Th25; A13: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 1) by A1,A5,A9,SCMPDS_2:67 .= inspos(0+(card I + 1)) by A3,Th23; s4.inspos(card I+1) = pIF.inspos(card I+1) by A10,A12,GRFUNC_1:8 .=halt SCMPDS by A11,Th25; then A14: CurInstr s4 = halt SCMPDS by A13,AMI_1:def 17; then A15: s3 is halting by AMI_1:def 20; now let l be Nat; assume l < 1; then l <1+0; then l <= 0 by NAT_1:38; then l=0 by NAT_1:18; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A4,Th29; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds 1 <= l; then LifeSpan s3 = 1 by A14,A15,SCM_1:def 2; then A16: s4 = Result s3 by A15,SCMFSA6B:16; A17: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (s +* SAl) by AMI_3:36; now let x be set; A18: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A19: x in dom IExec(IF,s); A20: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A19,SCMPDS_4:20; suppose A21: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A22: not x in dom SAl by A20,TARSKI:def 1; not x in dom (s | A) by A6,A21,SCMPDS_2:53; hence IExec(IF,s).x = s4.x by A16,A18,FUNCT_4:12 .= s3.x by A5,A21,SCMPDS_2:67 .= s.x by A21,SCMPDS_5:19 .= (s +* SAl).x by A22,FUNCT_4:12; suppose A23: x = IC SCMPDS; then x in dom Result s3 & not x in dom (s | A) by A6,AMI_1:48,AMI_5:25 ; hence IExec(IF,s).x = s4.x by A16,A18,FUNCT_4:12 .= inspos(card I + 1) by A13,A23,AMI_1:def 15 .= (s +* SAl).x by A23,Lm8; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (s +* SAl).x by Th27; end; hence IExec(IF,s) = s +* SAl by A17,FUNCT_1:9; end; definition let I be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if=0(a,k1,I) -> shiftable parahalting; correctness proof set IF=if=0(a,k1,I), IsIF=Initialized stop IF; set i = (a,k1)<>0_goto (card I +1); IF = i ';' I by Def7 .=Load i ';' I by SCMPDS_4:def 4; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1)= 0; then IF is_halting_on s by A2,Th62; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) <> 0; then IF is_halting_on s by Th63; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if=0(a,k1,I) -> No-StopCode; coherence proof if=0(a,k1,I) = (a,k1)<>0_goto (card I +1) ';' I by Def7; hence thesis; end; end; theorem ::E2,S8B_21A for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if=0(a,k1,I),s) = inspos (card I + 1) proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if=0(a,k1,I); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) = 0; then IExec(IF,s) =IExec(I,s) +* Start-At inspos (card I+1) by A1,Th64; hence thesis by AMI_5:79; suppose s.DataLoc(s.a,k1) <> 0; then IExec(IF,s) =s +* Start-At inspos (card I+ 1) by Th65; hence thesis by AMI_5:79; end; theorem ::E2,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1)= 0 holds IExec(if=0(a,k1,I),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a,b be Int_position,k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1)=0; then A2: IExec(if=0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 1) by A1,Th64; not b in dom Start-At inspos (card I + 1) by SCMPDS_4:59; hence IExec(if=0(a,k1,I),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::E2,S8B_21C for s being State of SCMPDS,I being Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1)<> 0 holds IExec(if=0(a,k1,I),s).b = s.b proof let s be State of SCMPDS,I be Program-block,a,b be Int_position, k1 be Integer; assume s.DataLoc(s.a,k1)<>0; then A1: IExec(if=0(a,k1,I),s) = s +* Start-At inspos (card I + 1) by Th65; not b in dom Start-At inspos (card I + 1) by SCMPDS_4:59; hence IExec(if=0(a,k1,I),s).b = s.b by A1,FUNCT_4:12; end; Lm9: card (i ';' j ';' I)=card I+2 proof thus card (i ';' j ';' I) =card (i ';' (j ';' I)) by SCMPDS_4:52 .=card (j ';' I)+1 by Th15 .=card I+1+1 by Th15 .=card I+(1+1) by XCMPLX_1:1 .=card I+2; end; begin :: The computation of "if var<>0 then block" theorem Th69: ::E3,S8B_14 card if<>0(a,k1,I) = card I + 2 proof if<>0(a,k1,I)=(a,k1)<>0_goto 2 ';' goto (card I +1) ';' I by Def8; hence thesis by Lm9; end; Lm10: inspos 0 in dom (i ';' j ';' I) & inspos 1 in dom (i ';' j ';' I) proof set ci=card (i ';' j ';' I); ci=card I + 2 by Lm9; then A1: 2 <= ci by NAT_1:29; then A2: 0 < ci by AXIOMS:22; 1 < ci by A1,AXIOMS:22; hence thesis by A2,SCMPDS_4:1; end; theorem Th70: ::LmT5 inspos 0 in dom if<>0(a,k1,I) & inspos 1 in dom if<>0(a,k1,I) proof if<>0(a,k1,I)=(a,k1)<>0_goto 2 ';' goto (card I +1) ';' I by Def8; hence thesis by Lm10; end; Lm11: (i ';' j ';' I).inspos 0=i & (i ';' j ';' I).inspos 1=j proof set jI=j ';' I; A1: i ';' j ';' I =i ';' jI by SCMPDS_4:52 .=Load i ';' jI by SCMPDS_4:def 4; inspos 0 in dom Load i by SCMPDS_5:2; hence (i ';' j ';' I).inspos 0 =(Load i).inspos 0 by A1,SCMPDS_4:37 .=i by SCMPDS_5:4; A2: card Load i=1 by SCMPDS_5:6; card jI=card I+1 by Th15; then 0 < card jI by NAT_1:19; then A3: inspos 0 in dom jI by SCMPDS_4:1; A4: inspos 0 in dom Load j by SCMPDS_5:2; thus (i ';' j ';' I).inspos 1 =(Load i ';' jI).inspos (0+1) by A1 .=(Load i ';' jI).(inspos 0+1) by SCMPDS_3:def 3 .=jI.inspos 0 by A2,A3,SCMPDS_4:38 .=(Load j ';' I).inspos 0 by SCMPDS_4:def 4 .=(Load j).inspos 0 by A4,SCMPDS_4:37 .=j by SCMPDS_5:4; end; theorem Th71: ::Lm6 if<>0(a,k1,I).inspos 0 = (a,k1)<>0_goto 2 & if<>0(a,k1,I).inspos 1 = goto (card I + 1) proof if<>0(a,k1,I)=(a,k1)<>0_goto 2 ';' goto (card I +1) ';' I by Def8; hence thesis by Lm11; end; theorem Th72: ::S8B_18 for s being State of SCMPDS, I being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)<>0 & I is_closed_on s & I is_halting_on s holds if<>0(a,k1,I) is_closed_on s & if<>0(a,k1,I) is_halting_on s proof let s be State of SCMPDS, I be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <> 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if<>0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<>0_goto 2, j = goto (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: inspos 0 in dom pIF by SCMPDS_4:75; A8: IF = i ';' j ';' I by Def8; then A9: IF = i ';' (j ';' I) by SCMPDS_4:52; A10: IC s3 =inspos 0 by Th21; A11: CurInstr s3 = i by A9,Th22; A12: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A11,AMI_1:def 18; A13: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A14: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A15: s3.DataLoc(s3.a,k1)=s3.b by A13,FUNCT_4:12 .= s.b by A14,FUNCT_4:12; A16: Shift(pI,2) c= s4 by A8,Lm7; A17: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,2) by A1,A12,A15,SCMPDS_2:67 .= inspos(0+2) by A10,Th23; s2,s3 equal_outside A by SCMPDS_4:36; then A18: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A18,SCMPDS_4:23 .= s4.a by A12,SCMPDS_2:67; end; then A19: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A16,A17,A19,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A21: k1 + 1 = k by NAT_1:22; consider m such that A22: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A23: card pIF = 1+ card IF by SCMPDS_5:7 .= 1+(card I +2) by Th69 .= 1+card I +2 by XCMPLX_1:1 .=card pI+2 by SCMPDS_5:7; inspos m in dom pI by A2,A22,Def2; then m < card pI by SCMPDS_4:1; then A24: m+2 < card pIF by A23,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A21,AMI_1:51 .= IC (Computation s2).k1 + 2 by A4,A6,A16,A17,A19,Th45 .= inspos (m + 2) by A22,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A24,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A7,A10,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A20,Def3; end; theorem Th73: ::E3,S8B_16 for s being State of SCMPDS,I being Program-block, a being Int_position, k1 being Integer st s.DataLoc(s.a,k1)= 0 holds if<>0(a,k1,I) is_closed_on s & if<>0(a,k1,I) is_halting_on s proof let s be State of SCMPDS,I be Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b = 0; set IF=if<>0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1, s5 = C3.2; set i = (a,k1)<>0_goto 2, j = goto (card I + 1); A2: IF = i ';' j ';' I by Def8 .=i ';' (j ';' I) by SCMPDS_4:52; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .=0 by A1,A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then A8: pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: inspos 1 in dom IF by Th70; then A11: inspos 1 in dom pIF by Th18; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A5,A7,SCMPDS_2:67 .= inspos(0+1) by A3,SCMPDS_4:70; s4.inspos 1 = pIF.inspos 1 by A9,A11,GRFUNC_1:8 .=IF.inspos 1 by A10,Th19 .=j by Th71; then A13: CurInstr s4 = j by A12,AMI_1:def 17; A14: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(j,s4) by A13,AMI_1:def 18; A15: pIF c= s5 by A8,AMI_3:38; A16: card IF=card I+2 by Th69; then A17: inspos(card I+2) in dom pIF by Th25; A18: IC s5 = s5.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card I+1) by A14,SCMPDS_2:66 .= inspos(card I+1+1) by A12,Th23 .= inspos(card I+(1+1)) by XCMPLX_1:1; s5.inspos(card I+2) = pIF.inspos(card I+2) by A15,A17,GRFUNC_1:8 .=halt SCMPDS by A16,Th25; then A19: CurInstr s5 = halt SCMPDS by A18,AMI_1:def 17; now let k be Nat; k = 0 or 0 < k by NAT_1:19; then A20: k = 0 or 0 + 1 <= k by INT_1:20; per cases by A20,REAL_1:def 5; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A3,SCMPDS_4:75; suppose k = 1; hence IC C3.k in dom pIF by A10,A12,Th18; suppose 1 < k; then 1+1 <= k by INT_1:20; hence IC C3.k in dom pIF by A17,A18,A19,AMI_1:52; end; hence IF is_closed_on s by Def2; s3 is halting by A19,AMI_1:def 20; hence IF is_halting_on s by Def3; end; theorem Th74: ::SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <> 0 & I is_closed_on s & I is_halting_on s holds IExec(if<>0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <> 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if<>0(a,k1,I), IsIF=Initialized stop IF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<>0_goto 2, j = goto (card I + 1); set SAl=Start-At inspos (card I + 2); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: IF = i ';' j ';' I by Def8; then A8: IF=i ';' (j ';' I) by SCMPDS_4:52; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: dom (s | A) = A by Th1; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: Shift(pI,2) c= s4 by A7,Lm7; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,2) by A1,A11,A14,SCMPDS_2:67 .= inspos(0+2) by A9,Th23; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:67; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; A19: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume A21: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A10,Th29; suppose l <> 0; then consider n such that A22: l = n + 1 by NAT_1:22; A23: n < LifeSpan s2 by A21,A22,AXIOMS:24; assume A24: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A22,A24,AMI_1:51; hence contradiction by A5,A23,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A25: LifeSpan s3 = LifeSpan s2 + 1 by A19,A20,SCM_1:def 2; A26: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A15,A16,A18,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A20,A25,SCMFSA6B:16; A27: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I,s) +* Start-At inspos (card I + 2)) by AMI_3:36; now let x be set; A28: IExec(I,s) = Result s2 +* s | A by SCMPDS_4:def 8; A29: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A30: x in dom IExec(IF,s); A31: dom Start-At inspos (card I + 2) = {IC SCMPDS} by AMI_3:34; per cases by A30,SCMPDS_4:20; suppose A32: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A33: not x in dom SAl by A31,TARSKI:def 1; A34: not x in dom (s | A) by A12,A32,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A29,FUNCT_4:12 .= (Result s2).x by A26,A32,SCMPDS_4:23 .= IExec(I,s).x by A28,A34,FUNCT_4:12 .= (IExec(I,s) +* SAl).x by A33,FUNCT_4:12; suppose A35: x = IC SCMPDS; then A36: x in dom SAl by A31,TARSKI:def 1; A37: not x in dom (s | A) by A12,A35,AMI_1:48; A38: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I,s).IC SCMPDS by A28,A35,A37,FUNCT_4:12 .= IC IExec(I,s) by AMI_1:def 15 .= inspos card I by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A29,A37,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A20,A25,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A35,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 2 by A4,A6,A15,A16,A18,Th45 .= IC Result s2 + 2 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I) + 2)).IC SCMPDS by A38,AMI_3:50 .= SAl.IC SCMPDS by SCMPDS_3:def 3 .= (IExec(I,s) +* SAl).x by A35,A36,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I,s) +* SAl by A27,FUNCT_1:9; end; theorem Th75: ::SCM8B_17 for s being State of SCMPDS,I being Program-block,a being Int_position, k1 being Integer st s.DataLoc(s.a,k1)= 0 holds IExec(if<>0(a,k1,I),s) = s +* Start-At inspos (card I + 2) proof let s be State of SCMPDS,I be Program-block,a be Int_position, k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b = 0; set IF=if<>0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1, s5 = C3.2; set i = (a,k1)<>0_goto 2, j = goto (card I + 1); set SAl=Start-At inspos (card I + 2); A2: IF = i ';' j ';' I by Def8 .=i ';' (j ';' I) by SCMPDS_4:52; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .=0 by A1,A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then A8: pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: inspos 1 in dom IF by Th70; then A11: inspos 1 in dom pIF by Th18; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A5,A7,SCMPDS_2:67 .= inspos(0+1) by A3,SCMPDS_4:70; s4.inspos 1 = pIF.inspos 1 by A9,A11,GRFUNC_1:8 .=IF.inspos 1 by A10,Th19 .=j by Th71; then A13: CurInstr s4 = j by A12,AMI_1:def 17; A14: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(j,s4) by A13,AMI_1:def 18; A15: pIF c= s5 by A8,AMI_3:38; A16: card IF=card I+2 by Th69; then A17: inspos(card I+2) in dom pIF by Th25; A18: IC s5 = s5.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card I+1) by A14,SCMPDS_2:66 .= inspos(card I+1+1) by A12,Th23 .= inspos(card I+(1+1)) by XCMPLX_1:1; s5.inspos(card I+2) = pIF.inspos(card I+2) by A15,A17,GRFUNC_1:8 .=halt SCMPDS by A16,Th25; then A19: CurInstr s5 = halt SCMPDS by A18,AMI_1:def 17; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume l < 1+1; then A21: l <= 1 by NAT_1:38; per cases by A21,CQC_THE1:2; suppose l=0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A4,Th29; suppose l=1; hence CurInstr (Computation s3).l <> halt SCMPDS by A13,SCMPDS_2:85; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds 2 <= l; then LifeSpan s3 = 2 by A19,A20,SCM_1:def 2; then A22: s5 = Result s3 by A20,SCMFSA6B:16; A23: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (s +* SAl) by AMI_3:36; A24: dom (s | A) = A by Th1; now let x be set; A25: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A26: x in dom IExec(IF,s); A27: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A26,SCMPDS_4:20; suppose A28: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A29: not x in dom SAl by A27,TARSKI:def 1; not x in dom (s | A) by A24,A28,SCMPDS_2:53; hence IExec(IF,s).x = s5.x by A22,A25,FUNCT_4:12 .= s4.x by A14,A28,SCMPDS_2:66 .= s3.x by A5,A28,SCMPDS_2:67 .= s.x by A28,SCMPDS_5:19 .= (s +* SAl).x by A29,FUNCT_4:12; suppose A30: x = IC SCMPDS; then x in dom Result s3 & not x in dom (s | A) by A24,AMI_1:48,AMI_5:25 ; hence IExec(IF,s).x = s5.x by A22,A25,FUNCT_4:12 .= inspos(card I + 2) by A18,A30,AMI_1:def 15 .= (s +* Start-At inspos (card I + 2)).x by A30,Lm8; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (s +* SAl).x by Th27; end; hence IExec(IF,s) = s +* SAl by A23,FUNCT_1:9; end; definition let I be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if<>0(a,k1,I) -> shiftable parahalting; correctness proof set IF=if<>0(a,k1,I), IsIF=Initialized stop IF; set i = (a,k1)<>0_goto 2, j = goto (card I + 1); IF = i ';' j ';' I by Def8 .=Load i ';' Load j ';' I by SCMPDS_4:def 6; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1)<>0; then IF is_halting_on s by A2,Th72; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) = 0; then IF is_halting_on s by Th73; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if<>0(a,k1,I) -> No-StopCode; coherence proof if<>0(a,k1,I) = (a,k1)<>0_goto 2 ';' goto (card I+1) ';' I by Def8; hence thesis; end; end; theorem ::E3,S8B_21A for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if<>0(a,k1,I),s) = inspos (card I + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if<>0(a,k1,I); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) <> 0; then IExec(IF,s) =IExec(I,s) +* Start-At inspos (card I+2) by A1,Th74; hence thesis by AMI_5:79; suppose s.DataLoc(s.a,k1) = 0; then IExec(IF,s) =s +* Start-At inspos (card I+ 2) by Th75; hence thesis by AMI_5:79; end; theorem ::E3,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1)<> 0 holds IExec(if<>0(a,k1,I),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a,b be Int_position,k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1)<>0; then A2: IExec(if<>0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 2) by A1,Th74; not b in dom Start-At inspos (card I + 2) by SCMPDS_4:59; hence IExec(if<>0(a,k1,I),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::E3,S8B_21C for s being State of SCMPDS,I being Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1)= 0 holds IExec(if<>0(a,k1,I),s).b = s.b proof let s be State of SCMPDS,I be Program-block,a,b be Int_position, k1 be Integer; assume s.DataLoc(s.a,k1)=0; then A1: IExec(if<>0(a,k1,I),s) = s +* Start-At inspos (card I + 2) by Th75; not b in dom Start-At inspos (card I + 2) by SCMPDS_4:59; hence IExec(if<>0(a,k1,I),s).b = s.b by A1,FUNCT_4:12; end; begin :: The computation of "if var>0 then block1 else block2" theorem Th79: ::G,S8B_14 card if>0(a,k1,I,J) = card I + card J + 2 proof if>0(a,k1,I,J)= (a,k1)<=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def5; hence thesis by Lm4; end; theorem inspos 0 in dom if>0(a,k1,I,J) & inspos 1 in dom if>0(a,k1,I,J) proof set ci=card if>0(a,k1,I,J); ci=card I + card J +2 by Th79; then 2 <= ci by NAT_1:37; then 0 < ci & 1 < ci by AXIOMS:22; hence thesis by SCMPDS_4:1; end; theorem if>0(a,k1,I,J).inspos 0 = (a,k1)<=0_goto (card I + 2) proof if>0(a,k1,I,J)= (a,k1)<=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def5; hence thesis by Lm5; end; theorem Th82: ::G,S8B_18 for s being State of SCMPDS, I,J being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)>0 & I is_closed_on s & I is_halting_on s holds if>0(a,k1,I,J) is_closed_on s & if>0(a,k1,I,J) is_halting_on s proof let s be State of SCMPDS, I,J be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b > 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set G=Goto (card J+1); set I2 = I ';' G ';' J, IF=if>0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, pI2=stop I2, II2= Initialized pI2, s2 = s +* II2, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<=0_goto (card I + 2); A4: II2 c= s2 by FUNCT_4:26; I2 is_halting_on s by A2,A3,Th44; then A5: s2 is halting by Def3; A6: I2 is_closed_on s by A2,A3,Th44; then A7: I2 is_closed_on s2 by Th38; A8: inspos 0 in dom pIF by SCMPDS_4:75; A9: IF = i ';' I ';' G ';' J by Def5 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' I2 by SCMPDS_4:50; A10: IC s3 =inspos 0 by Th21; A11: CurInstr s3 = i by A9,Th22; A12: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A11,AMI_1:def 18; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: card pIF = card IF +1 by SCMPDS_5:7 .= card I2 +1+1 by A9,Th15; A16: Shift(pI2,1) c= s4 by A9,Lm6; A17: I2 is shiftable by Lm3; A18: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A12,A14,SCMPDS_2:68 .= inspos(0+1) by A10,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A19: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A19,SCMPDS_4:23 .= s4.a by A12,SCMPDS_2:68; end; then A20: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A7,A16,A17,A18,A20,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A21: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A22: k1 + 1 = k by NAT_1:22; consider m such that A23: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A24: card pIF = card pI2+1 by A15,SCMPDS_5:7; inspos m in dom pI2 by A6,A23,Def2; then m < card pI2 by SCMPDS_4:1; then A25: m+1 < card pIF by A24,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A22,AMI_1:51 .= IC (Computation s2).k1 + 1 by A4,A7,A16,A17,A18,A20,Th45 .= inspos (m + 1) by A23,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A25,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A8,A10,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A21,Def3; end; theorem Th83: ::S8B_16 for s being State of SCMPDS,I being Program-block,J being shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <= 0 & J is_closed_on s & J is_halting_on s holds if>0(a,k1,I,J) is_closed_on s & if>0(a,k1,I,J) is_halting_on s proof let s be State of SCMPDS,I be Program-block,J be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <= 0; assume A2: J is_closed_on s; assume A3: J is_halting_on s; set pJ=stop J, IsJ=Initialized pJ, s1 = s +* IsJ, IF=if>0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<=0_goto (card I + 2); set G =Goto (card J+1), iG=i ';' I ';' G; A4: IsJ c= s1 by FUNCT_4:26; A5: s1 is halting by A3,Def3; A6: J is_closed_on s1 by A2,Th38; A7: IF = iG ';' J by Def5; then A8: IF = i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' (I ';' G ';' J) by SCMPDS_4:50; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: card iG = card (i ';' I) + card G by SCMPDS_4:45 .=card (i ';' I) + 1 by Th32 .=card I +1 +1 by Th15 .=card I +(1 +1) by XCMPLX_1:1; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: IsIF c= s3 by FUNCT_4:26; s1,s3 equal_outside A by SCMPDS_4:36; then A16: s1 | D = s3 | D by SCMPDS_4:24; now let a; thus s1.a = s3.a by A16,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:68; end; then A17: s1 | D = s4 | D by SCMPDS_4:23; A18: Shift(pJ,card I+2) c= pIF by A7,A12,Th24; pIF c= s3 by A15,SCMPDS_4:57; then Shift(pJ,card I+2) c= s3 by A18,XBOOLE_1:1; then A19: Shift(pJ,card I+2) c= s4 by AMI_3:38; A20: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 2) by A1,A11,A14,SCMPDS_2:68 .= inspos(0+(card I + 2)) by A9,Th23; CurInstr (Computation s3).(LifeSpan s1 + 1) =CurInstr (Computation s4).LifeSpan s1 by AMI_1:51 .=CurInstr (Computation s1).LifeSpan s1 by A4,A6,A17,A19,A20,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A21: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A22: k1 + 1 = k by NAT_1:22; consider m such that A23: inspos m = IC (Computation s1).k1 by SCMPDS_3:32; A24: card pJ = card J + 1 by SCMPDS_5:7; A25: card pIF = card IF+1 by SCMPDS_5:7 .=card I +2 +card J +1 by A7,A12,SCMPDS_4:45 .=card I +2 + card pJ by A24,XCMPLX_1:1; inspos m in dom pJ by A2,A23,Def2; then m < card pJ by SCMPDS_4:1; then A26: m + (card I + 2) < card pJ + (card I + 2) by REAL_1:53; IC C3.k = IC (Computation s4).k1 by A22,AMI_1:51 .= IC (Computation s1).k1 + (card I + 2) by A4,A6,A17,A19,A20,Th45 .= inspos (m + (card I + 2)) by A23,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A25,A26,SCMPDS_4:1; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A9,SCMPDS_4:75; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A21,Def3; end; theorem Th84: ::G,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, J being shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) > 0 & I is_closed_on s & I is_halting_on s holds IExec(if>0(a,k1,I,J),s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, J be shiftable Program-block,a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b > 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set G=Goto (card J+1); set I2 = I ';' G ';' J, IF=if>0(a,k1,I,J), IsIF=Initialized stop IF, pI2=stop I2, II2= Initialized pI2, s2 = s +* II2, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<=0_goto (card I + 2); set SAl= Start-At inspos (card I + card J + 2); A4: II2 c= s2 by FUNCT_4:26; I2 is_halting_on s by A2,A3,Th44; then A5: s2 is halting by Def3; I2 is_closed_on s by A2,A3,Th44; then A6: I2 is_closed_on s2 by Th38; A7: IF = i ';' I ';' G ';' J by Def5 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' I2 by SCMPDS_4:50; A8: IC s3 =inspos 0 by Th21; A9: CurInstr s3 = i by A7,Th22; A10: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A9,AMI_1:def 18; A11: dom (s | A) = A by Th1; A12: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A13: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A12,FUNCT_4:12; A14: Shift(pI2,1) c= s4 by A7,Lm6; A15: I2 is shiftable by Lm3; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A10,A13,SCMPDS_2:68 .= inspos(0+1) by A8,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A10,SCMPDS_2:68; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; A19: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A14,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume A21: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A9,Th30; suppose l <> 0; then consider n such that A22: l = n + 1 by NAT_1:22; A23: n < LifeSpan s2 by A21,A22,AXIOMS:24; assume A24: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A14,A15,A16,A18,Th45 .= halt SCMPDS by A22,A24,AMI_1:51; hence contradiction by A5,A23,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A25: LifeSpan s3 = LifeSpan s2 + 1 by A19,A20,SCM_1:def 2; A26: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A14,A15,A16,A18,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A20,A25,SCMFSA6B:16; A27: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I2,s) +* SAl) by AMI_3:36; now let x be set; A28: IExec(I2,s) = Result s2 +* s | A by SCMPDS_4:def 8; A29: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A30: x in dom IExec(IF,s); A31: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A30,SCMPDS_4:20; suppose A32: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A33: not x in dom SAl by A31,TARSKI:def 1; A34: not x in dom (s | A) by A11,A32,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A29,FUNCT_4:12 .= (Result s2).x by A26,A32,SCMPDS_4:23 .= IExec(I2,s).x by A28,A34,FUNCT_4:12 .= (IExec(I2,s) +* SAl).x by A33,FUNCT_4:12; suppose A35: x = IC SCMPDS; then A36: x in dom SAl by A31,TARSKI:def 1; A37: not x in dom (s | A) by A11,A35,AMI_1:48; A38: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I2,s).IC SCMPDS by A28,A35,A37,FUNCT_4:12 .= IC IExec(I2,s) by AMI_1:def 15 .= inspos (card I + card J + 1) by A2,A3,Th46; thus IExec(IF,s).x = (Result s3).x by A29,A37,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A20,A25,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A35,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 1 by A4,A6,A14,A15,A16,A18,Th45 .= IC Result s2 + 1 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I + card J + 1) + 1)).IC SCMPDS by A38,AMI_3:50 .= (Start-At inspos (card I + card J + 1 + 1)).IC SCMPDS by SCMPDS_3:def 3 .= (Start-At inspos (card I + card J + (1 + 1))).IC SCMPDS by XCMPLX_1:1 .= (IExec(I2,s) +* SAl).x by A35,A36,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I2,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I2,s) +* SAl by A27,FUNCT_1:9 .= IExec(I,s) +* Start-At inspos (card I + card J + 1) +* Start-At inspos (card I + card J + 2) by A2,A3,Th47 .= IExec(I,s) +* SAl by Th14; end; theorem Th85: ::G,SCM8B_17 for s being State of SCMPDS,I being Program-block,J being No-StopCode shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <= 0 & J is_closed_on s & J is_halting_on s holds IExec(if>0(a,k1,I,J),s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) proof let s be State of SCMPDS,I be Program-block,J be No-StopCode shiftable Program-block,a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <= 0; assume A2: J is_closed_on s; assume A3: J is_halting_on s; set pJ=stop J, IsJ=Initialized pJ, s1 = s +* IsJ, IF=if>0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<=0_goto (card I + 2); set G =Goto (card J+1), iG=i ';' I ';' G; set SAl=Start-At inspos (card I + card J + 2); A4: IsJ c= s1 by FUNCT_4:26; A5: s1 is halting by A3,Def3; A6: J is_closed_on s1 by A2,Th38; A7: IF = iG ';' J by Def5; then A8: IF = i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' (I ';' G ';' J) by SCMPDS_4:50; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: card iG = card (i ';' I) + card G by SCMPDS_4:45 .=card (i ';' I) + 1 by Th32 .=card I +1 +1 by Th15 .=card I +(1 +1) by XCMPLX_1:1; A13: dom (s | A) = A by Th1; A14: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A15: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A16: s3.DataLoc(s3.a,k1)=s3.b by A14,FUNCT_4:12 .= s.b by A15,FUNCT_4:12; A17: IsIF c= s3 by FUNCT_4:26; A18: Shift(pJ,card I+2) c= pIF by A7,A12,Th24; pIF c= s3 by A17,SCMPDS_4:57; then Shift(pJ,card I+2) c= s3 by A18,XBOOLE_1:1; then A19: Shift(pJ,card I+2) c= s4 by AMI_3:38; A20: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 2) by A1,A11,A16,SCMPDS_2:68 .= inspos(0+(card I + 2)) by A9,Th23; s1,s3 equal_outside A by SCMPDS_4:36; then A21: s1 | D = s3 | D by SCMPDS_4:24; now let a; thus s1.a = s3.a by A21,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:68; end; then A22: s1 | D = s4 | D by SCMPDS_4:23; A23: CurInstr (Computation s3).(LifeSpan s1 + 1) =CurInstr (Computation s4).LifeSpan s1 by AMI_1:51 .=CurInstr (Computation s1).LifeSpan s1 by A4,A6,A19,A20,A22,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A24: s3 is halting by AMI_1:def 20; now let l be Nat; assume A25: l < LifeSpan s1 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A10,Th30; suppose l <> 0; then consider n such that A26: l = n + 1 by NAT_1:22; A27: n < LifeSpan s1 by A25,A26,AXIOMS:24; assume A28: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s1).n = CurInstr (Computation s4).n by A4,A6,A19,A20,A22,Th45 .= halt SCMPDS by A26,A28,AMI_1:51; hence contradiction by A5,A27,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s1 + 1 <= l; then A29: LifeSpan s3 = LifeSpan s1 + 1 by A23,A24,SCM_1:def 2; A30: (Result s1) | D = (Computation s1).(LifeSpan s1) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s1) | D by A4,A6,A19,A20,A22,Th45 .= (Computation s3).(LifeSpan s1 + 1) | D by AMI_1:51 .= (Result s3) | D by A24,A29,SCMFSA6B:16; A31: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(J,s) +* SAl) by AMI_3:36; now let x be set; A32: IExec(J,s) = Result s1 +* s | A by SCMPDS_4:def 8; A33: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A34: x in dom IExec(IF,s); A35: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A34,SCMPDS_4:20; suppose A36: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A37: not x in dom SAl by A35,TARSKI:def 1; A38: not x in dom (s | A) by A13,A36,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A33,FUNCT_4:12 .= (Result s1).x by A30,A36,SCMPDS_4:23 .= IExec(J,s).x by A32,A38,FUNCT_4:12 .= (IExec(J,s) +* SAl).x by A37,FUNCT_4:12; suppose A39: x = IC SCMPDS; then A40: x in dom SAl by A35,TARSKI:def 1; A41: not x in dom (s | A) by A13,A39,AMI_1:48; A42: IC Result s1 = (Result s1).IC SCMPDS by AMI_1:def 15 .= IExec(J,s).IC SCMPDS by A32,A39,A41,FUNCT_4:12 .= IC IExec(J,s) by AMI_1:def 15 .= inspos (card J) by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A33,A41,FUNCT_4:12 .= (Computation s3).(LifeSpan s1 + 1).x by A24,A29,SCMFSA6B:16 .= (Computation s4).(LifeSpan s1).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s1) by A39,AMI_1:def 15 .= IC (Computation s1).(LifeSpan s1) + (card I + 2) by A4,A6,A19,A20,A22,Th45 .= IC Result s1 + (card I + 2) by A5,SCMFSA6B:16 .= (Start-At (inspos card J + (card I + 2))).IC SCMPDS by A42,AMI_3:50 .= (Start-At inspos (card I + 2+ card J)).IC SCMPDS by SCMPDS_3:def 3 .= SAl.IC SCMPDS by XCMPLX_1:1 .= (IExec(J,s) +* SAl).x by A39,A40,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(J,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(J,s) +* SAl by A31,FUNCT_1:9; end; definition let I,J be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if>0(a,k1,I,J) -> shiftable parahalting; correctness proof set IF=if>0(a,k1,I,J), IsIF=Initialized stop IF; set i = (a,k1)<=0_goto (card I + 2), G =Goto (card J+1); reconsider IJ=I ';' G ';' J as shiftable Program-block by Lm3; IF = i ';' I ';' G ';' J by Def5 .=i ';' (I ';' G) ';' J by SCMPDS_4:50 .=i ';' IJ by SCMPDS_4:50 .=Load i ';' IJ by SCMPDS_4:def 4; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; A3: J is_closed_on s & J is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) > 0; then IF is_halting_on s by A2,Th82; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) <= 0; then IF is_halting_on s by A3,Th83; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I,J be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if>0(a,k1,I,J) -> No-StopCode; coherence proof if>0(a,k1,I,J) = (a,k1)<=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def5; hence thesis; end; end; theorem ::G,S8B_21A for s being State of SCMPDS,I,J being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if>0(a,k1,I,J),s) = inspos (card I + card J + 2) proof let s be State of SCMPDS,I,J be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if>0(a,k1,I,J); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; A2: J is_closed_on s & J is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) > 0; then IExec(IF,s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) by A1,Th84; hence IC IExec(IF,s) = inspos (card I + card J + 2) by AMI_5:79; suppose s.DataLoc(s.a,k1) <= 0; then IExec(IF,s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) by A2,Th85; hence IC IExec(IF,s) = inspos (card I + card J + 2) by AMI_5:79; end; theorem ::G,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,J being shiftable Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1)>0 holds IExec(if>0(a,k1,I,J),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,J be shiftable Program-block,a,b be Int_position, k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1)>0; then A2: IExec(if>0(a,k1,I,J),s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) by A1,Th84; not b in dom Start-At inspos (card I + card J + 2) by SCMPDS_4:59; hence IExec(if>0(a,k1,I,J),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::G,S8B_21C for s being State of SCMPDS,I being Program-block,J being No-StopCode parahalting shiftable Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <= 0 holds IExec(if>0(a,k1,I,J),s).b = IExec(J,s).b proof let s be State of SCMPDS,I be Program-block,J be No-StopCode parahalting shiftable Program-block,a,b be Int_position, k1 be Integer; set IF=if>0(a,k1,I,J); A1: J is_closed_on s & J is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1) <= 0; then A2: IExec(IF,s) =IExec(J,s) +* Start-At inspos (card I + card J + 2) by A1,Th85; not b in dom Start-At inspos (card I + card J + 2) by SCMPDS_4:59; hence IExec(IF,s).b = IExec(J,s).b by A2,FUNCT_4:12; end; begin :: The computation of "if var>0 then block" theorem Th89: ::S8B_14 card if>0(a,k1,I) = card I + 1 proof thus card if>0(a,k1,I)=card ((a,k1)<=0_goto (card I +1) ';' I) by Def9 .=card I+1 by Th15; end; theorem ::LmT5 inspos 0 in dom if>0(a,k1,I) proof set ci=card if>0(a,k1,I); ci=card I + 1 by Th89; then 0 < ci by NAT_1:19; hence thesis by SCMPDS_4:1; end; theorem ::Lm6 if>0(a,k1,I).inspos 0 = (a,k1)<=0_goto (card I + 1) proof if>0(a,k1,I)=(a,k1)<=0_goto (card I +1) ';' I by Def9; hence thesis by Th16; end; theorem Th92: ::G2,S8B_18 for s being State of SCMPDS, I being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)> 0 & I is_closed_on s & I is_halting_on s holds if>0(a,k1,I) is_closed_on s & if>0(a,k1,I) is_halting_on s proof let s be State of SCMPDS, I be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b > 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if>0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<=0_goto (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: inspos 0 in dom pIF by SCMPDS_4:75; A8: IF = i ';' I by Def9; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A13: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A12,FUNCT_4:12; A14: card pIF = card IF +1 by SCMPDS_5:7 .= card I +1+1 by A8,Th15; A15: Shift(pI,1) c= s4 by A8,Lm6; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A11,A13,SCMPDS_2:68 .= inspos(0+1) by A9,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:68; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A19: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A20: k1 + 1 = k by NAT_1:22; consider m such that A21: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A22: card pIF = card pI+1 by A14,SCMPDS_5:7; inspos m in dom pI by A2,A21,Def2; then m < card pI by SCMPDS_4:1; then A23: m+1 < card pIF by A22,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A20,AMI_1:51 .= IC (Computation s2).k1 + 1 by A4,A6,A15,A16,A18,Th45 .= inspos (m + 1) by A21,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A23,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A7,A9,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A19,Def3; end; theorem Th93: ::G,S8B_16 for s being State of SCMPDS,I being Program-block, a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) <= 0 holds if>0(a,k1,I) is_closed_on s & if>0(a,k1,I) is_halting_on s proof let s be State of SCMPDS,I be Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <= 0; set IF=if>0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<=0_goto (card I + 1); A2: IF = i ';' I by Def9; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then pIF c= s3 by SCMPDS_4:57; then A8: pIF c= s4 by AMI_3:38; A9: card IF=card I+1 by Th89; then A10: inspos(card I+1) in dom pIF by Th25; A11: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 1) by A1,A5,A7,SCMPDS_2:68 .= inspos(0+(card I + 1)) by A3,Th23; s4.inspos(card I+1) = pIF.inspos(card I+1) by A8,A10,GRFUNC_1:8 .=halt SCMPDS by A9,Th25; then A12: CurInstr s4 = halt SCMPDS by A11,AMI_1:def 17; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then 1+0 <= k by INT_1:20; hence IC C3.k in dom pIF by A10,A11,A12,AMI_1:52; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A3,SCMPDS_4:75; end; hence IF is_closed_on s by Def2; s3 is halting by A12,AMI_1:def 20; hence IF is_halting_on s by Def3; end; theorem Th94: ::G2,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)> 0 & I is_closed_on s & I is_halting_on s holds IExec(if>0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 1) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b > 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if>0(a,k1,I), IsIF=Initialized stop IF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<=0_goto (card I + 1); set SAl=Start-At inspos (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: IF = i ';' I by Def9; A8: IC s3 =inspos 0 by Th21; A9: CurInstr s3 = i by A7,Th22; A10: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A9,AMI_1:def 18; A11: dom (s | A) = A by Th1; A12: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A13: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A12,FUNCT_4:12; A14: Shift(pI,1) c= s4 by A7,Lm6; A15: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A10,A13,SCMPDS_2:68 .= inspos(0+1) by A8,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A16: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A16,SCMPDS_4:23 .= s4.a by A10,SCMPDS_2:68; end; then A17: s2 | D = s4 | D by SCMPDS_4:23; A18: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A14,A15,A17,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A19: s3 is halting by AMI_1:def 20; now let l be Nat; assume A20: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A9,Th30; suppose l <> 0; then consider n such that A21: l = n + 1 by NAT_1:22; A22: n < LifeSpan s2 by A20,A21,AXIOMS:24; assume A23: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A14,A15,A17,Th45 .= halt SCMPDS by A21,A23,AMI_1:51; hence contradiction by A5,A22,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A24: LifeSpan s3 = LifeSpan s2 + 1 by A18,A19,SCM_1:def 2; A25: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A14,A15,A17,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A19,A24,SCMFSA6B:16; A26: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I,s) +* SAl) by AMI_3:36; now let x be set; A27: IExec(I,s) = Result s2 +* s | A by SCMPDS_4:def 8; A28: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A29: x in dom IExec(IF,s); A30: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A29,SCMPDS_4:20; suppose A31: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A32: not x in dom SAl by A30,TARSKI:def 1; A33: not x in dom (s | A) by A11,A31,SCMPDS_2:53; x in dom Result s3 & not x in dom (s | A) by A11,A31,SCMPDS_2:49,53; hence IExec(IF,s).x = (Result s3).x by A28,FUNCT_4:12 .= (Result s2).x by A25,A31,SCMPDS_4:23 .= IExec(I,s).x by A27,A33,FUNCT_4:12 .= (IExec(I,s) +* SAl).x by A32,FUNCT_4:12; suppose A34: x = IC SCMPDS; then A35: x in dom SAl by A30,TARSKI:def 1; A36: not x in dom (s | A) by A11,A34,AMI_1:48; A37: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I,s).IC SCMPDS by A27,A34,A36,FUNCT_4:12 .= IC IExec(I,s) by AMI_1:def 15 .= inspos card I by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A28,A36,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A19,A24,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A34,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 1 by A4,A6,A14,A15,A17,Th45 .= IC Result s2 + 1 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I) + 1)).IC SCMPDS by A37,AMI_3:50 .= SAl.IC SCMPDS by SCMPDS_3:def 3 .= (IExec(I,s) +* SAl).x by A34,A35,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I,s) +* SAl by A26,FUNCT_1:9; end; theorem Th95: ::G2,SCM8B_17 for s being State of SCMPDS,I being Program-block,a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) <= 0 holds IExec(if>0(a,k1,I),s) = s +* Start-At inspos (card I + 1) proof let s be State of SCMPDS,I be Program-block,a be Int_position, k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <= 0; set IF=if>0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<=0_goto (card I + 1); set SAl=Start-At inspos (card I + 1); A2: IF = i ';' I by Def9; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: dom (s | A) = A by Th1; A7: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A8: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A7,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: card IF=card I+1 by Th89; then A11: inspos(card I+1) in dom pIF by Th25; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 1) by A1,A5,A8,SCMPDS_2:68 .= inspos(0+(card I + 1)) by A3,Th23; s4.inspos(card I+1) = pIF.inspos(card I+1) by A9,A11,GRFUNC_1:8 .=halt SCMPDS by A10,Th25; then A13: CurInstr s4 = halt SCMPDS by A12,AMI_1:def 17; then A14: s3 is halting by AMI_1:def 20; now let l be Nat; assume l < 1; then l <1+0; then l <= 0 by NAT_1:38; then l=0 by NAT_1:18; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A4,Th30; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds 1 <= l; then LifeSpan s3 = 1 by A13,A14,SCM_1:def 2; then A15: s4 = Result s3 by A14,SCMFSA6B:16; A16: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (s +* SAl) by AMI_3:36; now let x be set; A17: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A18: x in dom IExec(IF,s); A19: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A18,SCMPDS_4:20; suppose A20: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A21: not x in dom SAl by A19,TARSKI:def 1; not x in dom (s | A) by A6,A20,SCMPDS_2:53; hence IExec(IF,s).x = s4.x by A15,A17,FUNCT_4:12 .= s3.x by A5,A20,SCMPDS_2:68 .= s.x by A20,SCMPDS_5:19 .= (s +* SAl).x by A21,FUNCT_4:12; suppose A22: x = IC SCMPDS; then x in dom Result s3 & not x in dom (s | A) by A6,AMI_1:48,AMI_5:25 ; hence IExec(IF,s).x = s4.x by A15,A17,FUNCT_4:12 .= inspos(card I + 1) by A12,A22,AMI_1:def 15 .= (s +* SAl).x by A22,Lm8; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (s +* SAl).x by Th27; end; hence IExec(IF,s) = s +* SAl by A16,FUNCT_1:9; end; definition let I be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if>0(a,k1,I) -> shiftable parahalting; correctness proof set IF=if>0(a,k1,I), IsIF=Initialized stop IF; set i = (a,k1)<=0_goto (card I +1); IF = i ';' I by Def9 .=Load i ';' I by SCMPDS_4:def 4; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) > 0; then IF is_halting_on s by A2,Th92; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) <= 0; then IF is_halting_on s by Th93; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if>0(a,k1,I) -> No-StopCode; coherence proof if>0(a,k1,I) = (a,k1)<=0_goto (card I +1) ';' I by Def9; hence thesis; end; end; theorem ::G,S8B_21A for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if>0(a,k1,I),s) = inspos (card I + 1) proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if>0(a,k1,I); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) > 0; then IExec(IF,s) =IExec(I,s) +* Start-At inspos (card I+1) by A1,Th94; hence thesis by AMI_5:79; suppose s.DataLoc(s.a,k1) <= 0; then IExec(IF,s) =s +* Start-At inspos (card I+ 1) by Th95; hence thesis by AMI_5:79; end; theorem ::G,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1)> 0 holds IExec(if>0(a,k1,I),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a,b be Int_position,k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1) > 0; then A2: IExec(if>0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 1) by A1,Th94; not b in dom Start-At inspos (card I + 1) by SCMPDS_4:59; hence IExec(if>0(a,k1,I),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::G,S8B_21C for s being State of SCMPDS,I being Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1) <= 0 holds IExec(if>0(a,k1,I),s).b = s.b proof let s be State of SCMPDS,I be Program-block,a,b be Int_position, k1 be Integer; assume s.DataLoc(s.a,k1) <= 0; then A1: IExec(if>0(a,k1,I),s) = s +* Start-At inspos (card I + 1) by Th95; not b in dom Start-At inspos (card I + 1) by SCMPDS_4:59; hence IExec(if>0(a,k1,I),s).b = s.b by A1,FUNCT_4:12; end; begin :: The computation of "if var<=0 then block" theorem Th99: ::S8B_14 card if<=0(a,k1,I) = card I + 2 proof if<=0(a,k1,I)=(a,k1)<=0_goto 2 ';' goto (card I +1) ';' I by Def10; hence thesis by Lm9; end; theorem Th100: ::LmT5 inspos 0 in dom if<=0(a,k1,I) & inspos 1 in dom if<=0(a,k1,I) proof if<=0(a,k1,I)=(a,k1)<=0_goto 2 ';' goto (card I +1) ';' I by Def10; hence thesis by Lm10; end; theorem Th101: ::Lm6 if<=0(a,k1,I).inspos 0 = (a,k1)<=0_goto 2 & if<=0(a,k1,I).inspos 1 = goto (card I + 1) proof if<=0(a,k1,I)=(a,k1)<=0_goto 2 ';' goto (card I +1) ';' I by Def10; hence thesis by Lm11; end; theorem Th102: ::S8B_18 for s being State of SCMPDS, I being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <= 0 & I is_closed_on s & I is_halting_on s holds if<=0(a,k1,I) is_closed_on s & if<=0(a,k1,I) is_halting_on s proof let s be State of SCMPDS, I be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <= 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if<=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)<=0_goto 2, j = goto (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: inspos 0 in dom pIF by SCMPDS_4:75; A8: IF = i ';' j ';' I by Def10; then A9: IF = i ';' (j ';' I) by SCMPDS_4:52; A10: IC s3 =inspos 0 by Th21; A11: CurInstr s3 = i by A9,Th22; A12: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A11,AMI_1:def 18; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: Shift(pI,2) c= s4 by A8,Lm7; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,2) by A1,A12,A14,SCMPDS_2:68 .= inspos(0+2) by A10,Th23; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A12,SCMPDS_2:68; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A19: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A20: k1 + 1 = k by NAT_1:22; consider m such that A21: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A22: card pIF = 1+ card IF by SCMPDS_5:7 .= 1+(card I +2) by Th99 .= 1+card I +2 by XCMPLX_1:1 .=card pI+2 by SCMPDS_5:7; inspos m in dom pI by A2,A21,Def2; then m < card pI by SCMPDS_4:1; then A23: m+2 < card pIF by A22,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A20,AMI_1:51 .= IC (Computation s2).k1 + 2 by A4,A6,A15,A16,A18,Th45 .= inspos (m + 2) by A21,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A23,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A7,A10,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A19,Def3; end; theorem Th103: ::G3,S8B_16 for s being State of SCMPDS,I being Program-block, a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) > 0 holds if<=0(a,k1,I) is_closed_on s & if<=0(a,k1,I) is_halting_on s proof let s be State of SCMPDS,I be Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b > 0; set IF=if<=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1, s5 = C3.2; set i = (a,k1)<=0_goto 2, j = goto (card I + 1); A2: IF = i ';' j ';' I by Def10 .=i ';' (j ';' I) by SCMPDS_4:52; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then A8: pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: inspos 1 in dom IF by Th100; then A11: inspos 1 in dom pIF by Th18; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A5,A7,SCMPDS_2:68 .= inspos(0+1) by A3,SCMPDS_4:70; s4.inspos 1 = pIF.inspos 1 by A9,A11,GRFUNC_1:8 .=IF.inspos 1 by A10,Th19 .=j by Th101; then A13: CurInstr s4 = j by A12,AMI_1:def 17; A14: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(j,s4) by A13,AMI_1:def 18; A15: pIF c= s5 by A8,AMI_3:38; A16: card IF=card I+2 by Th99; then A17: inspos(card I+2) in dom pIF by Th25; A18: IC s5 = s5.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card I+1) by A14,SCMPDS_2:66 .= inspos(card I+1+1) by A12,Th23 .= inspos(card I+(1+1)) by XCMPLX_1:1; s5.inspos(card I+2) = pIF.inspos(card I+2) by A15,A17,GRFUNC_1:8 .=halt SCMPDS by A16,Th25; then A19: CurInstr s5 = halt SCMPDS by A18,AMI_1:def 17; now let k be Nat; k = 0 or 0 < k by NAT_1:19; then A20: k = 0 or 0 + 1 <= k by INT_1:20; per cases by A20,REAL_1:def 5; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A3,SCMPDS_4:75; suppose k = 1; hence IC C3.k in dom pIF by A10,A12,Th18; suppose 1 < k; then 1+1 <= k by INT_1:20; hence IC C3.k in dom pIF by A17,A18,A19,AMI_1:52; end; hence IF is_closed_on s by Def2; s3 is halting by A19,AMI_1:def 20; hence IF is_halting_on s by Def3; end; theorem Th104: ::SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <= 0 & I is_closed_on s & I is_halting_on s holds IExec(if<=0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b <= 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if<=0(a,k1,I), IsIF=Initialized stop IF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)<=0_goto 2, j = goto (card I + 1); set SAl=Start-At inspos (card I + 2); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: IF = i ';' j ';' I by Def10; then A8: IF=i ';' (j ';' I) by SCMPDS_4:52; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: dom (s | A) = A by Th1; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: Shift(pI,2) c= s4 by A7,Lm7; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,2) by A1,A11,A14,SCMPDS_2:68 .= inspos(0+2) by A9,Th23; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:68; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; A19: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume A21: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A10,Th30; suppose l <> 0; then consider n such that A22: l = n + 1 by NAT_1:22; A23: n < LifeSpan s2 by A21,A22,AXIOMS:24; assume A24: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A22,A24,AMI_1:51; hence contradiction by A5,A23,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A25: LifeSpan s3 = LifeSpan s2 + 1 by A19,A20,SCM_1:def 2; A26: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A15,A16,A18,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A20,A25,SCMFSA6B:16; A27: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I,s) +* Start-At inspos (card I + 2)) by AMI_3:36; now let x be set; A28: IExec(I,s) = Result s2 +* s | A by SCMPDS_4:def 8; A29: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A30: x in dom IExec(IF,s); A31: dom Start-At inspos (card I + 2) = {IC SCMPDS} by AMI_3:34; per cases by A30,SCMPDS_4:20; suppose A32: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A33: not x in dom SAl by A31,TARSKI:def 1; A34: not x in dom (s | A) by A12,A32,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A29,FUNCT_4:12 .= (Result s2).x by A26,A32,SCMPDS_4:23 .= IExec(I,s).x by A28,A34,FUNCT_4:12 .= (IExec(I,s) +* SAl).x by A33,FUNCT_4:12; suppose A35: x = IC SCMPDS; then A36: x in dom SAl by A31,TARSKI:def 1; A37: not x in dom (s | A) by A12,A35,AMI_1:48; A38: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I,s).IC SCMPDS by A28,A35,A37,FUNCT_4:12 .= IC IExec(I,s) by AMI_1:def 15 .= inspos card I by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A29,A37,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A20,A25,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A35,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 2 by A4,A6,A15,A16,A18,Th45 .= IC Result s2 + 2 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I) + 2)).IC SCMPDS by A38,AMI_3:50 .= SAl.IC SCMPDS by SCMPDS_3:def 3 .= (IExec(I,s) +* SAl).x by A35,A36,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I,s) +* SAl by A27,FUNCT_1:9; end; theorem Th105: ::G3,SCM8B_17 for s being State of SCMPDS,I being Program-block,a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) > 0 holds IExec(if<=0(a,k1,I),s) = s +* Start-At inspos (card I + 2) proof let s be State of SCMPDS,I be Program-block,a be Int_position, k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b > 0; set IF=if<=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1, s5 = C3.2; set i = (a,k1)<=0_goto 2, j = goto (card I + 1); set SAl=Start-At inspos (card I + 2); A2: IF = i ';' j ';' I by Def10 .=i ';' (j ';' I) by SCMPDS_4:52; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then A8: pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: inspos 1 in dom IF by Th100; then A11: inspos 1 in dom pIF by Th18; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A5,A7,SCMPDS_2:68 .= inspos(0+1) by A3,SCMPDS_4:70; s4.inspos 1 = pIF.inspos 1 by A9,A11,GRFUNC_1:8 .=IF.inspos 1 by A10,Th19 .=j by Th101; then A13: CurInstr s4 = j by A12,AMI_1:def 17; A14: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(j,s4) by A13,AMI_1:def 18; A15: pIF c= s5 by A8,AMI_3:38; A16: card IF=card I+2 by Th99; then A17: inspos(card I+2) in dom pIF by Th25; A18: IC s5 = s5.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card I+1) by A14,SCMPDS_2:66 .= inspos(card I+1+1) by A12,Th23 .= inspos(card I+(1+1)) by XCMPLX_1:1; s5.inspos(card I+2) = pIF.inspos(card I+2) by A15,A17,GRFUNC_1:8 .=halt SCMPDS by A16,Th25; then A19: CurInstr s5 = halt SCMPDS by A18,AMI_1:def 17; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume l < 1+1; then A21: l <= 1 by NAT_1:38; per cases by A21,CQC_THE1:2; suppose l=0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A4,Th30; suppose l=1; hence CurInstr (Computation s3).l <> halt SCMPDS by A13,SCMPDS_2:85; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds 2 <= l; then LifeSpan s3 = 2 by A19,A20,SCM_1:def 2; then A22: s5 = Result s3 by A20,SCMFSA6B:16; A23: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (s +* SAl) by AMI_3:36; A24: dom (s | A) = A by Th1; now let x be set; A25: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A26: x in dom IExec(IF,s); A27: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A26,SCMPDS_4:20; suppose A28: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A29: not x in dom SAl by A27,TARSKI:def 1; not x in dom (s | A) by A24,A28,SCMPDS_2:53; hence IExec(IF,s).x = s5.x by A22,A25,FUNCT_4:12 .= s4.x by A14,A28,SCMPDS_2:66 .= s3.x by A5,A28,SCMPDS_2:68 .= s.x by A28,SCMPDS_5:19 .= (s +* SAl).x by A29,FUNCT_4:12; suppose A30: x = IC SCMPDS; then x in dom Result s3 & not x in dom (s | A) by A24,AMI_1:48,AMI_5:25 ; hence IExec(IF,s).x = s5.x by A22,A25,FUNCT_4:12 .= inspos(card I + 2) by A18,A30,AMI_1:def 15 .= (s +* Start-At inspos (card I + 2)).x by A30,Lm8; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (s +* SAl).x by Th27; end; hence IExec(IF,s) = s +* SAl by A23,FUNCT_1:9; end; definition let I be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if<=0(a,k1,I) -> shiftable parahalting; correctness proof set IF=if<=0(a,k1,I), IsIF=Initialized stop IF; set i = (a,k1)<=0_goto 2, j = goto (card I + 1); IF = i ';' j ';' I by Def10 .=Load i ';' Load j ';' I by SCMPDS_4:def 6; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) <= 0; then IF is_halting_on s by A2,Th102; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) > 0; then IF is_halting_on s by Th103; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if<=0(a,k1,I) -> No-StopCode; coherence proof if<=0(a,k1,I) = (a,k1)<=0_goto 2 ';' goto (card I+1) ';' I by Def10; hence thesis; end; end; theorem ::G3,S8B_21A for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if<=0(a,k1,I),s) = inspos (card I + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if<=0(a,k1,I); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) <= 0; then IExec(IF,s) =IExec(I,s) +* Start-At inspos (card I+2) by A1,Th104; hence thesis by AMI_5:79; suppose s.DataLoc(s.a,k1) > 0; then IExec(IF,s) =s +* Start-At inspos (card I+ 2) by Th105; hence thesis by AMI_5:79; end; theorem ::G3,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1) <= 0 holds IExec(if<=0(a,k1,I),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a,b be Int_position,k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1) <= 0; then A2: IExec(if<=0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 2) by A1,Th104; not b in dom Start-At inspos (card I + 2) by SCMPDS_4:59; hence IExec(if<=0(a,k1,I),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::S8B_21C for s being State of SCMPDS,I being Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1) > 0 holds IExec(if<=0(a,k1,I),s).b = s.b proof let s be State of SCMPDS,I be Program-block,a,b be Int_position, k1 be Integer; assume s.DataLoc(s.a,k1) > 0; then A1: IExec(if<=0(a,k1,I),s) = s +* Start-At inspos (card I + 2) by Th105; not b in dom Start-At inspos (card I + 2) by SCMPDS_4:59; hence IExec(if<=0(a,k1,I),s).b = s.b by A1,FUNCT_4:12; end; begin :: The computation of "if var<0 then block1 else block2" theorem Th109: ::L,S8B_14 card if<0(a,k1,I,J) = card I + card J + 2 proof if<0(a,k1,I,J)= (a,k1)>=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def6; hence thesis by Lm4; end; theorem inspos 0 in dom if<0(a,k1,I,J) & inspos 1 in dom if<0(a,k1,I,J) proof set ci=card if<0(a,k1,I,J); ci=card I + card J +2 by Th109; then 2 <= ci by NAT_1:37; then 0 < ci & 1 < ci by AXIOMS:22; hence thesis by SCMPDS_4:1; end; theorem if<0(a,k1,I,J).inspos 0 = (a,k1)>=0_goto (card I + 2) proof if<0(a,k1,I,J)= (a,k1)>=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def6; hence thesis by Lm5; end; theorem Th112: ::L,S8B_18 for s being State of SCMPDS, I,J being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1)<0 & I is_closed_on s & I is_halting_on s holds if<0(a,k1,I,J) is_closed_on s & if<0(a,k1,I,J) is_halting_on s proof let s be State of SCMPDS, I,J be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b < 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set G=Goto (card J+1); set I2 = I ';' G ';' J, IF=if<0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, pI2=stop I2, II2= Initialized pI2, s2 = s +* II2, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)>=0_goto (card I + 2); A4: II2 c= s2 by FUNCT_4:26; I2 is_halting_on s by A2,A3,Th44; then A5: s2 is halting by Def3; A6: I2 is_closed_on s by A2,A3,Th44; then A7: I2 is_closed_on s2 by Th38; A8: inspos 0 in dom pIF by SCMPDS_4:75; A9: IF = i ';' I ';' G ';' J by Def6 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' I2 by SCMPDS_4:50; A10: IC s3 =inspos 0 by Th21; A11: CurInstr s3 = i by A9,Th22; A12: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A11,AMI_1:def 18; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: card pIF = card IF +1 by SCMPDS_5:7 .= card I2 +1+1 by A9,Th15; A16: Shift(pI2,1) c= s4 by A9,Lm6; A17: I2 is shiftable by Lm3; A18: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A12,A14,SCMPDS_2:69 .= inspos(0+1) by A10,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A19: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A19,SCMPDS_4:23 .= s4.a by A12,SCMPDS_2:69; end; then A20: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A7,A16,A17,A18,A20,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A21: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A22: k1 + 1 = k by NAT_1:22; consider m such that A23: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A24: card pIF = card pI2+1 by A15,SCMPDS_5:7; inspos m in dom pI2 by A6,A23,Def2; then m < card pI2 by SCMPDS_4:1; then A25: m+1 < card pIF by A24,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A22,AMI_1:51 .= IC (Computation s2).k1 + 1 by A4,A7,A16,A17,A18,A20,Th45 .= inspos (m + 1) by A23,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A25,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A8,A10,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A21,Def3; end; theorem Th113: ::L,S8B_16 for s being State of SCMPDS,I being Program-block,J being shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) >= 0 & J is_closed_on s & J is_halting_on s holds if<0(a,k1,I,J) is_closed_on s & if<0(a,k1,I,J) is_halting_on s proof let s be State of SCMPDS,I be Program-block,J be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b >= 0; assume A2: J is_closed_on s; assume A3: J is_halting_on s; set pJ=stop J, IsJ=Initialized pJ, s1 = s +* IsJ, IF=if<0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)>=0_goto (card I + 2); set G =Goto (card J+1), iG=i ';' I ';' G; A4: IsJ c= s1 by FUNCT_4:26; A5: s1 is halting by A3,Def3; A6: J is_closed_on s1 by A2,Th38; A7: IF = iG ';' J by Def6; then A8: IF = i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' (I ';' G ';' J) by SCMPDS_4:50; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: card iG = card (i ';' I) + card G by SCMPDS_4:45 .=card (i ';' I) + 1 by Th32 .=card I +1 +1 by Th15 .=card I +(1 +1) by XCMPLX_1:1; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: IsIF c= s3 by FUNCT_4:26; s1,s3 equal_outside A by SCMPDS_4:36; then A16: s1 | D = s3 | D by SCMPDS_4:24; now let a; thus s1.a = s3.a by A16,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:69; end; then A17: s1 | D = s4 | D by SCMPDS_4:23; A18: Shift(pJ,card I+2) c= pIF by A7,A12,Th24; pIF c= s3 by A15,SCMPDS_4:57; then Shift(pJ,card I+2) c= s3 by A18,XBOOLE_1:1; then A19: Shift(pJ,card I+2) c= s4 by AMI_3:38; A20: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 2) by A1,A11,A14,SCMPDS_2:69 .= inspos(0+(card I + 2)) by A9,Th23; CurInstr (Computation s3).(LifeSpan s1 + 1) =CurInstr (Computation s4).LifeSpan s1 by AMI_1:51 .=CurInstr (Computation s1).LifeSpan s1 by A4,A6,A17,A19,A20,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A21: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A22: k1 + 1 = k by NAT_1:22; consider m such that A23: inspos m = IC (Computation s1).k1 by SCMPDS_3:32; A24: card pJ = card J + 1 by SCMPDS_5:7; A25: card pIF = card IF+1 by SCMPDS_5:7 .=card I +2 +card J +1 by A7,A12,SCMPDS_4:45 .=card I +2 + card pJ by A24,XCMPLX_1:1; inspos m in dom pJ by A2,A23,Def2; then m < card pJ by SCMPDS_4:1; then A26: m + (card I + 2) < card pJ + (card I + 2) by REAL_1:53; IC C3.k = IC (Computation s4).k1 by A22,AMI_1:51 .= IC (Computation s1).k1 + (card I + 2) by A4,A6,A17,A19,A20,Th45 .= inspos (m + (card I + 2)) by A23,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A25,A26,SCMPDS_4:1; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A9,SCMPDS_4:75; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A21,Def3; end; theorem Th114: ::L,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, J being shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) < 0 & I is_closed_on s & I is_halting_on s holds IExec(if<0(a,k1,I,J),s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, J be shiftable Program-block,a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b < 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set G=Goto (card J+1); set I2 = I ';' G ';' J, IF=if<0(a,k1,I,J), IsIF=Initialized stop IF, pI2=stop I2, II2= Initialized pI2, s2 = s +* II2, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)>=0_goto (card I + 2); set SAl= Start-At inspos (card I + card J + 2); A4: II2 c= s2 by FUNCT_4:26; I2 is_halting_on s by A2,A3,Th44; then A5: s2 is halting by Def3; I2 is_closed_on s by A2,A3,Th44; then A6: I2 is_closed_on s2 by Th38; A7: IF = i ';' I ';' G ';' J by Def6 .= i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' I2 by SCMPDS_4:50; A8: IC s3 =inspos 0 by Th21; A9: CurInstr s3 = i by A7,Th22; A10: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A9,AMI_1:def 18; A11: dom (s | A) = A by Th1; A12: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A13: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A12,FUNCT_4:12; A14: Shift(pI2,1) c= s4 by A7,Lm6; A15: I2 is shiftable by Lm3; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A10,A13,SCMPDS_2:69 .= inspos(0+1) by A8,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A10,SCMPDS_2:69; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; A19: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A14,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume A21: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A9,Th31; suppose l <> 0; then consider n such that A22: l = n + 1 by NAT_1:22; A23: n < LifeSpan s2 by A21,A22,AXIOMS:24; assume A24: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A14,A15,A16,A18,Th45 .= halt SCMPDS by A22,A24,AMI_1:51; hence contradiction by A5,A23,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A25: LifeSpan s3 = LifeSpan s2 + 1 by A19,A20,SCM_1:def 2; A26: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A14,A15,A16,A18,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A20,A25,SCMFSA6B:16; A27: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I2,s) +* SAl) by AMI_3:36; now let x be set; A28: IExec(I2,s) = Result s2 +* s | A by SCMPDS_4:def 8; A29: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A30: x in dom IExec(IF,s); A31: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A30,SCMPDS_4:20; suppose A32: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A33: not x in dom SAl by A31,TARSKI:def 1; A34: not x in dom (s | A) by A11,A32,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A29,FUNCT_4:12 .= (Result s2).x by A26,A32,SCMPDS_4:23 .= IExec(I2,s).x by A28,A34,FUNCT_4:12 .= (IExec(I2,s) +* SAl).x by A33,FUNCT_4:12; suppose A35: x = IC SCMPDS; then A36: x in dom SAl by A31,TARSKI:def 1; A37: not x in dom (s | A) by A11,A35,AMI_1:48; A38: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I2,s).IC SCMPDS by A28,A35,A37,FUNCT_4:12 .= IC IExec(I2,s) by AMI_1:def 15 .= inspos (card I + card J + 1) by A2,A3,Th46; thus IExec(IF,s).x = (Result s3).x by A29,A37,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A20,A25,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A35,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 1 by A4,A6,A14,A15,A16,A18,Th45 .= IC Result s2 + 1 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I + card J + 1) + 1)).IC SCMPDS by A38,AMI_3:50 .= (Start-At inspos (card I + card J + 1 + 1)).IC SCMPDS by SCMPDS_3:def 3 .= (Start-At inspos (card I + card J + (1 + 1))).IC SCMPDS by XCMPLX_1:1 .= (IExec(I2,s) +* SAl).x by A35,A36,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I2,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I2,s) +* SAl by A27,FUNCT_1:9 .= IExec(I,s) +* Start-At inspos (card I + card J + 1) +* Start-At inspos (card I + card J + 2) by A2,A3,Th47 .= IExec(I,s) +* SAl by Th14; end; theorem Th115: ::L,SCM8B_17 for s being State of SCMPDS,I being Program-block,J being No-StopCode shiftable Program-block,a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) >= 0 & J is_closed_on s & J is_halting_on s holds IExec(if<0(a,k1,I,J),s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) proof let s be State of SCMPDS,I be Program-block,J be No-StopCode shiftable Program-block,a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b >= 0; assume A2: J is_closed_on s; assume A3: J is_halting_on s; set pJ=stop J, IsJ=Initialized pJ, s1 = s +* IsJ, IF=if<0(a,k1,I,J), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)>=0_goto (card I + 2); set G =Goto (card J+1), iG=i ';' I ';' G; set SAl=Start-At inspos (card I + card J + 2); A4: IsJ c= s1 by FUNCT_4:26; A5: s1 is halting by A3,Def3; A6: J is_closed_on s1 by A2,Th38; A7: IF = iG ';' J by Def6; then A8: IF = i ';' (I ';' G) ';' J by SCMPDS_4:50 .= i ';' (I ';' G ';' J) by SCMPDS_4:50; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: card iG = card (i ';' I) + card G by SCMPDS_4:45 .=card (i ';' I) + 1 by Th32 .=card I +1 +1 by Th15 .=card I +(1 +1) by XCMPLX_1:1; A13: dom (s | A) = A by Th1; A14: not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; A15: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; A16: s3.DataLoc(s3.a,k1)=s3.b by A14,FUNCT_4:12 .= s.b by A15,FUNCT_4:12; A17: IsIF c= s3 by FUNCT_4:26; A18: Shift(pJ,card I+2) c= pIF by A7,A12,Th24; pIF c= s3 by A17,SCMPDS_4:57; then Shift(pJ,card I+2) c= s3 by A18,XBOOLE_1:1; then A19: Shift(pJ,card I+2) c= s4 by AMI_3:38; A20: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 2) by A1,A11,A16,SCMPDS_2:69 .= inspos(0+(card I + 2)) by A9,Th23; s1,s3 equal_outside A by SCMPDS_4:36; then A21: s1 | D = s3 | D by SCMPDS_4:24; now let a; thus s1.a = s3.a by A21,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:69; end; then A22: s1 | D = s4 | D by SCMPDS_4:23; A23: CurInstr (Computation s3).(LifeSpan s1 + 1) =CurInstr (Computation s4).LifeSpan s1 by AMI_1:51 .=CurInstr (Computation s1).LifeSpan s1 by A4,A6,A19,A20,A22,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A24: s3 is halting by AMI_1:def 20; now let l be Nat; assume A25: l < LifeSpan s1 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A10,Th31; suppose l <> 0; then consider n such that A26: l = n + 1 by NAT_1:22; A27: n < LifeSpan s1 by A25,A26,AXIOMS:24; assume A28: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s1).n = CurInstr (Computation s4).n by A4,A6,A19,A20,A22,Th45 .= halt SCMPDS by A26,A28,AMI_1:51; hence contradiction by A5,A27,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s1 + 1 <= l; then A29: LifeSpan s3 = LifeSpan s1 + 1 by A23,A24,SCM_1:def 2; A30: (Result s1) | D = (Computation s1).(LifeSpan s1) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s1) | D by A4,A6,A19,A20,A22,Th45 .= (Computation s3).(LifeSpan s1 + 1) | D by AMI_1:51 .= (Result s3) | D by A24,A29,SCMFSA6B:16; A31: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(J,s) +* SAl) by AMI_3:36; now let x be set; A32: IExec(J,s) = Result s1 +* s | A by SCMPDS_4:def 8; A33: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A34: x in dom IExec(IF,s); A35: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A34,SCMPDS_4:20; suppose A36: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A37: not x in dom SAl by A35,TARSKI:def 1; A38: not x in dom (s | A) by A13,A36,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A33,FUNCT_4:12 .= (Result s1).x by A30,A36,SCMPDS_4:23 .= IExec(J,s).x by A32,A38,FUNCT_4:12 .= (IExec(J,s) +* SAl).x by A37,FUNCT_4:12; suppose A39: x = IC SCMPDS; then A40: x in dom SAl by A35,TARSKI:def 1; A41: not x in dom (s | A) by A13,A39,AMI_1:48; A42: IC Result s1 = (Result s1).IC SCMPDS by AMI_1:def 15 .= IExec(J,s).IC SCMPDS by A32,A39,A41,FUNCT_4:12 .= IC IExec(J,s) by AMI_1:def 15 .= inspos (card J) by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A33,A41,FUNCT_4:12 .= (Computation s3).(LifeSpan s1 + 1).x by A24,A29,SCMFSA6B:16 .= (Computation s4).(LifeSpan s1).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s1) by A39,AMI_1:def 15 .= IC (Computation s1).(LifeSpan s1) + (card I + 2) by A4,A6,A19,A20,A22,Th45 .= IC Result s1 + (card I + 2) by A5,SCMFSA6B:16 .= (Start-At (inspos card J + (card I + 2))).IC SCMPDS by A42,AMI_3:50 .= (Start-At inspos (card I + 2+ card J)).IC SCMPDS by SCMPDS_3:def 3 .= SAl.IC SCMPDS by XCMPLX_1:1 .= (IExec(J,s) +* SAl).x by A39,A40,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(J,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(J,s) +* SAl by A31,FUNCT_1:9; end; definition let I,J be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if<0(a,k1,I,J) -> shiftable parahalting; correctness proof set IF=if<0(a,k1,I,J), IsIF=Initialized stop IF; set i = (a,k1)>=0_goto (card I + 2), G =Goto (card J+1); reconsider IJ=I ';' G ';' J as shiftable Program-block by Lm3; IF = i ';' I ';' G ';' J by Def6 .=i ';' (I ';' G) ';' J by SCMPDS_4:50 .=i ';' IJ by SCMPDS_4:50 .=Load i ';' IJ by SCMPDS_4:def 4; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; A3: J is_closed_on s & J is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) < 0; then IF is_halting_on s by A2,Th112; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) >= 0; then IF is_halting_on s by A3,Th113; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I,J be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if<0(a,k1,I,J) -> No-StopCode; coherence proof if<0(a,k1,I,J) = (a,k1)>=0_goto (card I +2) ';' I ';' Goto (card J+1) ';' J by Def6; hence thesis; end; end; theorem ::L,S8B_21A for s being State of SCMPDS,I,J being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if<0(a,k1,I,J),s) = inspos (card I + card J + 2) proof let s be State of SCMPDS,I,J be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if<0(a,k1,I,J); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; A2: J is_closed_on s & J is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) < 0; then IExec(IF,s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) by A1,Th114; hence IC IExec(IF,s) = inspos (card I + card J + 2) by AMI_5:79; suppose s.DataLoc(s.a,k1) >= 0; then IExec(IF,s) = IExec(J,s) +* Start-At inspos (card I + card J + 2) by A2,Th115; hence IC IExec(IF,s) = inspos (card I + card J + 2) by AMI_5:79; end; theorem ::S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,J being shiftable Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1)<0 holds IExec(if<0(a,k1,I,J),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,J be shiftable Program-block,a,b be Int_position, k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1)<0; then A2: IExec(if<0(a,k1,I,J),s) = IExec(I,s) +* Start-At inspos (card I + card J + 2) by A1,Th114; not b in dom Start-At inspos (card I + card J + 2) by SCMPDS_4:59; hence IExec(if<0(a,k1,I,J),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::L,S8B_21C for s being State of SCMPDS,I being Program-block,J being No-StopCode parahalting shiftable Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1) >= 0 holds IExec(if<0(a,k1,I,J),s).b = IExec(J,s).b proof let s be State of SCMPDS,I be Program-block,J be No-StopCode parahalting shiftable Program-block,a,b be Int_position, k1 be Integer; set IF=if<0(a,k1,I,J); A1: J is_closed_on s & J is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1) >= 0; then A2: IExec(IF,s) =IExec(J,s) +* Start-At inspos (card I + card J + 2) by A1,Th115; not b in dom Start-At inspos (card I + card J + 2) by SCMPDS_4:59; hence IExec(IF,s).b = IExec(J,s).b by A2,FUNCT_4:12; end; begin :: The computation of "if var<0 then block" theorem Th119: ::L2,S8B_14 card if<0(a,k1,I) = card I + 1 proof thus card if<0(a,k1,I)=card ((a,k1)>=0_goto (card I +1) ';' I) by Def11 .=card I+1 by Th15; end; theorem ::LmT5 inspos 0 in dom if<0(a,k1,I) proof set ci=card if<0(a,k1,I); ci=card I + 1 by Th119; then 0 < ci by NAT_1:19; hence thesis by SCMPDS_4:1; end; theorem ::Lm6 if<0(a,k1,I).inspos 0 = (a,k1)>=0_goto (card I + 1) proof if<0(a,k1,I)=(a,k1)>=0_goto (card I +1) ';' I by Def11; hence thesis by Th16; end; theorem Th122: ::L2,S8B_18 for s being State of SCMPDS, I being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) < 0 & I is_closed_on s & I is_halting_on s holds if<0(a,k1,I) is_closed_on s & if<0(a,k1,I) is_halting_on s proof let s be State of SCMPDS, I be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b < 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if<0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)>=0_goto (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: inspos 0 in dom pIF by SCMPDS_4:75; A8: IF = i ';' I by Def11; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A13: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A12,FUNCT_4:12; A14: card pIF = card IF +1 by SCMPDS_5:7 .= card I +1+1 by A8,Th15; A15: Shift(pI,1) c= s4 by A8,Lm6; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A11,A13,SCMPDS_2:69 .= inspos(0+1) by A9,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:69; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A19: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A20: k1 + 1 = k by NAT_1:22; consider m such that A21: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A22: card pIF = card pI+1 by A14,SCMPDS_5:7; inspos m in dom pI by A2,A21,Def2; then m < card pI by SCMPDS_4:1; then A23: m+1 < card pIF by A22,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A20,AMI_1:51 .= IC (Computation s2).k1 + 1 by A4,A6,A15,A16,A18,Th45 .= inspos (m + 1) by A21,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A23,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A7,A9,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A19,Def3; end; theorem Th123: ::L,S8B_16 for s being State of SCMPDS,I being Program-block, a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) >= 0 holds if<0(a,k1,I) is_closed_on s & if<0(a,k1,I) is_halting_on s proof let s be State of SCMPDS,I be Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b >= 0; set IF=if<0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)>=0_goto (card I + 1); A2: IF = i ';' I by Def11; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then pIF c= s3 by SCMPDS_4:57; then A8: pIF c= s4 by AMI_3:38; A9: card IF=card I+1 by Th119; then A10: inspos(card I+1) in dom pIF by Th25; A11: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 1) by A1,A5,A7,SCMPDS_2:69 .= inspos(0+(card I + 1)) by A3,Th23; s4.inspos(card I+1) = pIF.inspos(card I+1) by A8,A10,GRFUNC_1:8 .=halt SCMPDS by A9,Th25; then A12: CurInstr s4 = halt SCMPDS by A11,AMI_1:def 17; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then 1+0 <= k by INT_1:20; hence IC C3.k in dom pIF by A10,A11,A12,AMI_1:52; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A3,SCMPDS_4:75; end; hence IF is_closed_on s by Def2; s3 is halting by A12,AMI_1:def 20; hence IF is_halting_on s by Def3; end; theorem Th124: ::L,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) < 0 & I is_closed_on s & I is_halting_on s holds IExec(if<0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 1) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b < 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if<0(a,k1,I), IsIF=Initialized stop IF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)>=0_goto (card I + 1); set SAl=Start-At inspos (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: IF = i ';' I by Def11; A8: IC s3 =inspos 0 by Th21; A9: CurInstr s3 = i by A7,Th22; A10: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A9,AMI_1:def 18; A11: dom (s | A) = A by Th1; A12: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A13: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A12,FUNCT_4:12; A14: Shift(pI,1) c= s4 by A7,Lm6; A15: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A10,A13,SCMPDS_2:69 .= inspos(0+1) by A8,SCMPDS_4:70; s2,s3 equal_outside A by SCMPDS_4:36; then A16: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A16,SCMPDS_4:23 .= s4.a by A10,SCMPDS_2:69; end; then A17: s2 | D = s4 | D by SCMPDS_4:23; A18: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A14,A15,A17,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A19: s3 is halting by AMI_1:def 20; now let l be Nat; assume A20: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A9,Th31; suppose l <> 0; then consider n such that A21: l = n + 1 by NAT_1:22; A22: n < LifeSpan s2 by A20,A21,AXIOMS:24; assume A23: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A14,A15,A17,Th45 .= halt SCMPDS by A21,A23,AMI_1:51; hence contradiction by A5,A22,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A24: LifeSpan s3 = LifeSpan s2 + 1 by A18,A19,SCM_1:def 2; A25: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A14,A15,A17,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A19,A24,SCMFSA6B:16; A26: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I,s) +* SAl) by AMI_3:36; now let x be set; A27: IExec(I,s) = Result s2 +* s | A by SCMPDS_4:def 8; A28: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A29: x in dom IExec(IF,s); A30: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A29,SCMPDS_4:20; suppose A31: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A32: not x in dom SAl by A30,TARSKI:def 1; A33: not x in dom (s | A) by A11,A31,SCMPDS_2:53; x in dom Result s3 & not x in dom (s | A) by A11,A31,SCMPDS_2:49,53; hence IExec(IF,s).x = (Result s3).x by A28,FUNCT_4:12 .= (Result s2).x by A25,A31,SCMPDS_4:23 .= IExec(I,s).x by A27,A33,FUNCT_4:12 .= (IExec(I,s) +* SAl).x by A32,FUNCT_4:12; suppose A34: x = IC SCMPDS; then A35: x in dom SAl by A30,TARSKI:def 1; A36: not x in dom (s | A) by A11,A34,AMI_1:48; A37: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I,s).IC SCMPDS by A27,A34,A36,FUNCT_4:12 .= IC IExec(I,s) by AMI_1:def 15 .= inspos card I by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A28,A36,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A19,A24,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A34,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 1 by A4,A6,A14,A15,A17,Th45 .= IC Result s2 + 1 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I) + 1)).IC SCMPDS by A37,AMI_3:50 .= SAl.IC SCMPDS by SCMPDS_3:def 3 .= (IExec(I,s) +* SAl).x by A34,A35,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I,s) +* SAl by A26,FUNCT_1:9; end; theorem Th125: ::L,SCM8B_17 for s being State of SCMPDS,I being Program-block,a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) >= 0 holds IExec(if<0(a,k1,I),s) = s +* Start-At inspos (card I + 1) proof let s be State of SCMPDS,I be Program-block,a be Int_position, k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b >= 0; set IF=if<0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)>=0_goto (card I + 1); set SAl=Start-At inspos (card I + 1); A2: IF = i ';' I by Def11; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: dom (s | A) = A by Th1; A7: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A8: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A7,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: card IF=card I+1 by Th119; then A11: inspos(card I+1) in dom pIF by Th25; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,card I + 1) by A1,A5,A8,SCMPDS_2:69 .= inspos(0+(card I + 1)) by A3,Th23; s4.inspos(card I+1) = pIF.inspos(card I+1) by A9,A11,GRFUNC_1:8 .=halt SCMPDS by A10,Th25; then A13: CurInstr s4 = halt SCMPDS by A12,AMI_1:def 17; then A14: s3 is halting by AMI_1:def 20; now let l be Nat; assume l < 1; then l <1+0; then l <= 0 by NAT_1:38; then l=0 by NAT_1:18; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A4,Th31; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds 1 <= l; then LifeSpan s3 = 1 by A13,A14,SCM_1:def 2; then A15: s4 = Result s3 by A14,SCMFSA6B:16; A16: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (s +* SAl) by AMI_3:36; now let x be set; A17: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A18: x in dom IExec(IF,s); A19: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A18,SCMPDS_4:20; suppose A20: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A21: not x in dom SAl by A19,TARSKI:def 1; not x in dom (s | A) by A6,A20,SCMPDS_2:53; hence IExec(IF,s).x = s4.x by A15,A17,FUNCT_4:12 .= s3.x by A5,A20,SCMPDS_2:69 .= s.x by A20,SCMPDS_5:19 .= (s +* SAl).x by A21,FUNCT_4:12; suppose A22: x = IC SCMPDS; then x in dom Result s3 & not x in dom (s | A) by A6,AMI_1:48,AMI_5:25 ; hence IExec(IF,s).x = s4.x by A15,A17,FUNCT_4:12 .= inspos(card I + 1) by A12,A22,AMI_1:def 15 .= (s +* SAl).x by A22,Lm8; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (s +* SAl).x by Th27; end; hence IExec(IF,s) = s +* SAl by A16,FUNCT_1:9; end; definition let I be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if<0(a,k1,I) -> shiftable parahalting; correctness proof set IF=if<0(a,k1,I), IsIF=Initialized stop IF; set i = (a,k1)>=0_goto (card I +1); IF = i ';' I by Def11 .=Load i ';' I by SCMPDS_4:def 4; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) < 0; then IF is_halting_on s by A2,Th122; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) >= 0; then IF is_halting_on s by Th123; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if<0(a,k1,I) -> No-StopCode; coherence proof if<0(a,k1,I) = (a,k1)>=0_goto (card I +1) ';' I by Def11; hence thesis; end; end; theorem ::L2,S8B_21A for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if<0(a,k1,I),s) = inspos (card I + 1) proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if<0(a,k1,I); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) < 0; then IExec(IF,s) =IExec(I,s) +* Start-At inspos (card I+1) by A1,Th124; hence thesis by AMI_5:79; suppose s.DataLoc(s.a,k1) >= 0; then IExec(IF,s) =s +* Start-At inspos (card I+ 1) by Th125; hence thesis by AMI_5:79; end; theorem ::L,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1) < 0 holds IExec(if<0(a,k1,I),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a,b be Int_position,k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1) < 0; then A2: IExec(if<0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 1) by A1,Th124; not b in dom Start-At inspos (card I + 1) by SCMPDS_4:59; hence IExec(if<0(a,k1,I),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::L2,S8B_21C for s being State of SCMPDS,I being Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1) >= 0 holds IExec(if<0(a,k1,I),s).b = s.b proof let s be State of SCMPDS,I be Program-block,a,b be Int_position, k1 be Integer; assume s.DataLoc(s.a,k1) >= 0; then A1: IExec(if<0(a,k1,I),s) = s +* Start-At inspos (card I + 1) by Th125; not b in dom Start-At inspos (card I + 1) by SCMPDS_4:59; hence IExec(if<0(a,k1,I),s).b = s.b by A1,FUNCT_4:12; end; begin :: The computation of "if var>=0 then block" theorem Th129: ::L3,S8B_14 card if>=0(a,k1,I) = card I + 2 proof if>=0(a,k1,I)=(a,k1)>=0_goto 2 ';' goto (card I +1) ';' I by Def12; hence thesis by Lm9; end; theorem Th130: ::LmT5 inspos 0 in dom if>=0(a,k1,I) & inspos 1 in dom if>=0(a,k1,I) proof if>=0(a,k1,I)=(a,k1)>=0_goto 2 ';' goto (card I +1) ';' I by Def12; hence thesis by Lm10; end; theorem Th131: ::Lm6 if>=0(a,k1,I).inspos 0 = (a,k1)>=0_goto 2 & if>=0(a,k1,I).inspos 1 = goto (card I + 1) proof if>=0(a,k1,I)=(a,k1)>=0_goto 2 ';' goto (card I +1) ';' I by Def12; hence thesis by Lm11; end; theorem Th132: ::S8B_18 for s being State of SCMPDS, I being shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) >= 0 & I is_closed_on s & I is_halting_on s holds if>=0(a,k1,I) is_closed_on s & if>=0(a,k1,I) is_halting_on s proof let s be State of SCMPDS, I be shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b >= 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if>=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1; set i = (a,k1)>=0_goto 2, j = goto (card I + 1); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: inspos 0 in dom pIF by SCMPDS_4:75; A8: IF = i ';' j ';' I by Def12; then A9: IF = i ';' (j ';' I) by SCMPDS_4:52; A10: IC s3 =inspos 0 by Th21; A11: CurInstr s3 = i by A9,Th22; A12: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A11,AMI_1:def 18; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: Shift(pI,2) c= s4 by A8,Lm7; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,2) by A1,A12,A14,SCMPDS_2:69 .= inspos(0+2) by A10,Th23; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A12,SCMPDS_2:69; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A19: s3 is halting by AMI_1:def 20; now let k be Nat; per cases by NAT_1:19; suppose 0 < k; then consider k1 being Nat such that A20: k1 + 1 = k by NAT_1:22; consider m such that A21: inspos m = IC (Computation s2).k1 by SCMPDS_3:32; A22: card pIF = 1+ card IF by SCMPDS_5:7 .= 1+(card I +2) by Th129 .= 1+card I +2 by XCMPLX_1:1 .=card pI+2 by SCMPDS_5:7; inspos m in dom pI by A2,A21,Def2; then m < card pI by SCMPDS_4:1; then A23: m+2 < card pIF by A22,REAL_1:53; IC C3.k = IC (Computation s4).k1 by A20,AMI_1:51 .= IC (Computation s2).k1 + 2 by A4,A6,A15,A16,A18,Th45 .= inspos (m + 2) by A21,SCMPDS_3:def 3; hence IC C3.k in dom pIF by A23,SCMPDS_4:1; suppose k = 0; hence IC C3.k in dom pIF by A7,A10,AMI_1:def 19; end; hence IF is_closed_on s by Def2; thus IF is_halting_on s by A19,Def3; end; theorem Th133: ::L3,S8B_16 for s being State of SCMPDS,I being Program-block, a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) < 0 holds if>=0(a,k1,I) is_closed_on s & if>=0(a,k1,I) is_halting_on s proof let s be State of SCMPDS,I be Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b < 0; set IF=if>=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1, s5 = C3.2; set i = (a,k1)>=0_goto 2, j = goto (card I + 1); A2: IF = i ';' j ';' I by Def12 .=i ';' (j ';' I) by SCMPDS_4:52; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then A8: pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: inspos 1 in dom IF by Th130; then A11: inspos 1 in dom pIF by Th18; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A5,A7,SCMPDS_2:69 .= inspos(0+1) by A3,SCMPDS_4:70; s4.inspos 1 = pIF.inspos 1 by A9,A11,GRFUNC_1:8 .=IF.inspos 1 by A10,Th19 .=j by Th131; then A13: CurInstr s4 = j by A12,AMI_1:def 17; A14: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(j,s4) by A13,AMI_1:def 18; A15: pIF c= s5 by A8,AMI_3:38; A16: card IF=card I+2 by Th129; then A17: inspos(card I+2) in dom pIF by Th25; A18: IC s5 = s5.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card I+1) by A14,SCMPDS_2:66 .= inspos(card I+1+1) by A12,Th23 .= inspos(card I+(1+1)) by XCMPLX_1:1; s5.inspos(card I+2) = pIF.inspos(card I+2) by A15,A17,GRFUNC_1:8 .=halt SCMPDS by A16,Th25; then A19: CurInstr s5 = halt SCMPDS by A18,AMI_1:def 17; now let k be Nat; k = 0 or 0 < k by NAT_1:19; then A20: k = 0 or 0 + 1 <= k by INT_1:20; per cases by A20,REAL_1:def 5; suppose k = 0; then C3.k = s3 by AMI_1:def 19; hence IC C3.k in dom pIF by A3,SCMPDS_4:75; suppose k = 1; hence IC C3.k in dom pIF by A10,A12,Th18; suppose 1 < k; then 1+1 <= k by INT_1:20; hence IC C3.k in dom pIF by A17,A18,A19,AMI_1:52; end; hence IF is_closed_on s by Def2; s3 is halting by A19,AMI_1:def 20; hence IF is_halting_on s by Def3; end; theorem Th134: ::L,SCM8B_19 for s being State of SCMPDS,I being No-StopCode shiftable Program-block, a being Int_position,k1 being Integer st s.DataLoc(s.a,k1) >= 0 & I is_closed_on s & I is_halting_on s holds IExec(if>=0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable Program-block, a be Int_position,k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b >= 0; assume A2: I is_closed_on s; assume A3: I is_halting_on s; set IF=if>=0(a,k1,I), IsIF=Initialized stop IF, pI=stop I, IsI= Initialized pI, s2 = s +* IsI, s3 = s +* IsIF, s4 = (Computation s3).1; set i = (a,k1)>=0_goto 2, j = goto (card I + 1); set SAl=Start-At inspos (card I + 2); A4: IsI c= s2 by FUNCT_4:26; A5: s2 is halting by A3,Def3; A6: I is_closed_on s2 by A2,Th38; A7: IF = i ';' j ';' I by Def12; then A8: IF=i ';' (j ';' I) by SCMPDS_4:52; A9: IC s3 =inspos 0 by Th21; A10: CurInstr s3 = i by A8,Th22; A11: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A10,AMI_1:def 18; A12: dom (s | A) = A by Th1; A13: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A14: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A13,FUNCT_4:12; A15: Shift(pI,2) c= s4 by A7,Lm7; A16: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s3,2) by A1,A11,A14,SCMPDS_2:69 .= inspos(0+2) by A9,Th23; s2,s3 equal_outside A by SCMPDS_4:36; then A17: s2 | D = s3 | D by SCMPDS_4:24; now let a; thus s2.a = s3.a by A17,SCMPDS_4:23 .= s4.a by A11,SCMPDS_2:69; end; then A18: s2 | D = s4 | D by SCMPDS_4:23; A19: CurInstr (Computation s3).(LifeSpan s2 + 1) =CurInstr (Computation s4).LifeSpan s2 by AMI_1:51 .=CurInstr (Computation s2).LifeSpan s2 by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A5,SCM_1:def 2; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume A21: l < LifeSpan s2 + 1; per cases; suppose l = 0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A10,Th31; suppose l <> 0; then consider n such that A22: l = n + 1 by NAT_1:22; A23: n < LifeSpan s2 by A21,A22,AXIOMS:24; assume A24: CurInstr (Computation s3).l = halt SCMPDS; CurInstr (Computation s2).n = CurInstr (Computation s4).n by A4,A6,A15,A16,A18,Th45 .= halt SCMPDS by A22,A24,AMI_1:51; hence contradiction by A5,A23,SCM_1:def 2; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds LifeSpan s2 + 1 <= l; then A25: LifeSpan s3 = LifeSpan s2 + 1 by A19,A20,SCM_1:def 2; A26: (Result s2) | D = (Computation s2).(LifeSpan s2) | D by A5,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2) | D by A4,A6,A15,A16,A18,Th45 .= (Computation s3).(LifeSpan s2 + 1) | D by AMI_1:51 .= (Result s3) | D by A20,A25,SCMFSA6B:16; A27: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (IExec(I,s) +* Start-At inspos (card I + 2)) by AMI_3:36; now let x be set; A28: IExec(I,s) = Result s2 +* s | A by SCMPDS_4:def 8; A29: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A30: x in dom IExec(IF,s); A31: dom Start-At inspos (card I + 2) = {IC SCMPDS} by AMI_3:34; per cases by A30,SCMPDS_4:20; suppose A32: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A33: not x in dom SAl by A31,TARSKI:def 1; A34: not x in dom (s | A) by A12,A32,SCMPDS_2:53; hence IExec(IF,s).x = (Result s3).x by A29,FUNCT_4:12 .= (Result s2).x by A26,A32,SCMPDS_4:23 .= IExec(I,s).x by A28,A34,FUNCT_4:12 .= (IExec(I,s) +* SAl).x by A33,FUNCT_4:12; suppose A35: x = IC SCMPDS; then A36: x in dom SAl by A31,TARSKI:def 1; A37: not x in dom (s | A) by A12,A35,AMI_1:48; A38: IC Result s2 = (Result s2).IC SCMPDS by AMI_1:def 15 .= IExec(I,s).IC SCMPDS by A28,A35,A37,FUNCT_4:12 .= IC IExec(I,s) by AMI_1:def 15 .= inspos card I by A2,A3,Th48; thus IExec(IF,s).x = (Result s3).x by A29,A37,FUNCT_4:12 .= (Computation s3).(LifeSpan s2 + 1).x by A20,A25,SCMFSA6B:16 .= (Computation s4).(LifeSpan s2).x by AMI_1:51 .= IC (Computation s4).(LifeSpan s2) by A35,AMI_1:def 15 .= IC (Computation s2).(LifeSpan s2) + 2 by A4,A6,A15,A16,A18,Th45 .= IC Result s2 + 2 by A5,SCMFSA6B:16 .= (Start-At (inspos (card I) + 2)).IC SCMPDS by A38,AMI_3:50 .= SAl.IC SCMPDS by SCMPDS_3:def 3 .= (IExec(I,s) +* SAl).x by A35,A36,FUNCT_4:14; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (IExec(I,s) +* SAl).x by Th26; end; hence IExec(IF,s) = IExec(I,s) +* SAl by A27,FUNCT_1:9; end; theorem Th135: ::L,SCM8B_17 for s being State of SCMPDS,I being Program-block,a being Int_position, k1 being Integer st s.DataLoc(s.a,k1) < 0 holds IExec(if>=0(a,k1,I),s) = s +* Start-At inspos (card I + 2) proof let s be State of SCMPDS,I be Program-block,a be Int_position, k1 be Integer; set b=DataLoc(s.a,k1); assume A1: s.b < 0; set IF=if>=0(a,k1,I), pIF=stop IF, IsIF=Initialized pIF, s3 = s +* IsIF, C3 = Computation s3, s4 = C3.1, s5 = C3.2; set i = (a,k1)>=0_goto 2, j = goto (card I + 1); set SAl=Start-At inspos (card I + 2); A2: IF = i ';' j ';' I by Def12 .=i ';' (j ';' I) by SCMPDS_4:52; A3: IC s3 =inspos 0 by Th21; A4: CurInstr s3 = i by A2,Th22; A5: (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A4,AMI_1:def 18; A6: not b in dom IsIF & b in dom s by SCMPDS_2:49,SCMPDS_4:31; not a in dom IsIF & a in dom s by SCMPDS_2:49,SCMPDS_4:31; then A7: s3.DataLoc(s3.a,k1)=s3.b by FUNCT_4:12 .= s.b by A6,FUNCT_4:12; IsIF c= s3 by FUNCT_4:26; then A8: pIF c= s3 by SCMPDS_4:57; then A9: pIF c= s4 by AMI_3:38; A10: inspos 1 in dom IF by Th130; then A11: inspos 1 in dom pIF by Th18; A12: IC s4 = s4.IC SCMPDS by AMI_1:def 15 .= Next IC s3 by A1,A5,A7,SCMPDS_2:69 .= inspos(0+1) by A3,SCMPDS_4:70; s4.inspos 1 = pIF.inspos 1 by A9,A11,GRFUNC_1:8 .=IF.inspos 1 by A10,Th19 .=j by Th131; then A13: CurInstr s4 = j by A12,AMI_1:def 17; A14: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(j,s4) by A13,AMI_1:def 18; A15: pIF c= s5 by A8,AMI_3:38; A16: card IF=card I+2 by Th129; then A17: inspos(card I+2) in dom pIF by Th25; A18: IC s5 = s5.IC SCMPDS by AMI_1:def 15 .= ICplusConst(s4,card I+1) by A14,SCMPDS_2:66 .= inspos(card I+1+1) by A12,Th23 .= inspos(card I+(1+1)) by XCMPLX_1:1; s5.inspos(card I+2) = pIF.inspos(card I+2) by A15,A17,GRFUNC_1:8 .=halt SCMPDS by A16,Th25; then A19: CurInstr s5 = halt SCMPDS by A18,AMI_1:def 17; then A20: s3 is halting by AMI_1:def 20; now let l be Nat; assume l < 1+1; then A21: l <= 1 by NAT_1:38; per cases by A21,CQC_THE1:2; suppose l=0; then CurInstr (Computation s3).l = CurInstr s3 by AMI_1:def 19; hence CurInstr (Computation s3).l <> halt SCMPDS by A4,Th31; suppose l=1; hence CurInstr (Computation s3).l <> halt SCMPDS by A13,SCMPDS_2:85; end; then for l be Nat st CurInstr (Computation s3).l = halt SCMPDS holds 2 <= l; then LifeSpan s3 = 2 by A19,A20,SCM_1:def 2; then A22: s5 = Result s3 by A20,SCMFSA6B:16; A23: dom IExec(IF,s) = the carrier of SCMPDS by AMI_3:36 .= dom (s +* SAl) by AMI_3:36; A24: dom (s | A) = A by Th1; now let x be set; A25: IExec(IF,s) = Result s3 +* s | A by SCMPDS_4:def 8; assume A26: x in dom IExec(IF,s); A27: dom SAl = {IC SCMPDS} by AMI_3:34; per cases by A26,SCMPDS_4:20; suppose A28: x is Int_position; then x <> IC SCMPDS by SCMPDS_2:52; then A29: not x in dom SAl by A27,TARSKI:def 1; not x in dom (s | A) by A24,A28,SCMPDS_2:53; hence IExec(IF,s).x = s5.x by A22,A25,FUNCT_4:12 .= s4.x by A14,A28,SCMPDS_2:66 .= s3.x by A5,A28,SCMPDS_2:69 .= s.x by A28,SCMPDS_5:19 .= (s +* SAl).x by A29,FUNCT_4:12; suppose A30: x = IC SCMPDS; then x in dom Result s3 & not x in dom (s | A) by A24,AMI_1:48,AMI_5:25 ; hence IExec(IF,s).x = s5.x by A22,A25,FUNCT_4:12 .= inspos(card I + 2) by A18,A30,AMI_1:def 15 .= (s +* Start-At inspos (card I + 2)).x by A30,Lm8; suppose x is Instruction-Location of SCMPDS; hence IExec(IF,s).x = (s +* SAl).x by Th27; end; hence IExec(IF,s) = s +* SAl by A23,FUNCT_1:9; end; definition let I be shiftable parahalting Program-block, a be Int_position,k1 be Integer; cluster if>=0(a,k1,I) -> shiftable parahalting; correctness proof set IF=if>=0(a,k1,I), IsIF=Initialized stop IF; set i = (a,k1)>=0_goto 2, j = goto (card I + 1); IF = i ';' j ';' I by Def12 .=Load i ';' Load j ';' I by SCMPDS_4:def 6; hence IF is shiftable; now let s be State of SCMPDS; assume IsIF c= s; then A1: s = s +* IsIF by AMI_5:10; A2: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) >= 0; then IF is_halting_on s by A2,Th132; hence s is halting by A1,Def3; suppose s.DataLoc(s.a,k1) < 0; then IF is_halting_on s by Th133; hence s is halting by A1,Def3; end; then IsIF is halting by AMI_1:def 26; hence IF is parahalting by SCMPDS_4:def 10; end; end; definition let I be No-StopCode Program-block, a be Int_position,k1 be Integer; cluster if>=0(a,k1,I) -> No-StopCode; coherence proof if>=0(a,k1,I) = (a,k1)>=0_goto 2 ';' goto (card I+1) ';' I by Def12; hence thesis; end; end; theorem ::L2,S8B_21A for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a being Int_position,k1 being Integer holds IC IExec(if>=0(a,k1,I),s) = inspos (card I + 2) proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a be Int_position,k1 be Integer; set IF=if>=0(a,k1,I); A1: I is_closed_on s & I is_halting_on s by Th34,Th35; per cases; suppose s.DataLoc(s.a,k1) >= 0; then IExec(IF,s) =IExec(I,s) +* Start-At inspos (card I+2) by A1,Th134; hence thesis by AMI_5:79; suppose s.DataLoc(s.a,k1) < 0; then IExec(IF,s) =s +* Start-At inspos (card I+ 2) by Th135; hence thesis by AMI_5:79; end; theorem ::L,S8B_21B for s being State of SCMPDS,I being No-StopCode shiftable parahalting Program-block,a,b being Int_position,k1 being Integer st s.DataLoc(s.a,k1) >= 0 holds IExec(if>=0(a,k1,I),s).b = IExec(I,s).b proof let s be State of SCMPDS,I be No-StopCode shiftable parahalting Program-block,a,b be Int_position,k1 be Integer; A1: I is_closed_on s & I is_halting_on s by Th34,Th35; assume s.DataLoc(s.a,k1) >= 0; then A2: IExec(if>=0(a,k1,I),s) = IExec(I,s) +* Start-At inspos (card I + 2) by A1,Th134; not b in dom Start-At inspos (card I + 2) by SCMPDS_4:59; hence IExec(if>=0(a,k1,I),s).b = IExec(I,s).b by A2,FUNCT_4:12; end; theorem ::L,S8B_21C for s being State of SCMPDS,I being Program-block,a,b being Int_position, k1 being Integer st s.DataLoc(s.a,k1) < 0 holds IExec(if>=0(a,k1,I),s).b = s.b proof let s be State of SCMPDS,I be Program-block,a,b be Int_position, k1 be Integer; assume s.DataLoc(s.a,k1) < 0; then A1: IExec(if>=0(a,k1,I),s) = s +* Start-At inspos (card I + 2) by Th135; not b in dom Start-At inspos (card I + 2) by SCMPDS_4:59; hence IExec(if>=0(a,k1,I),s).b = s.b by A1,FUNCT_4:12; end;