Copyright (c) 1997 Association of Mizar Users
environ vocabulary AMI_1, SCMFSA_2, BOOLE, SCMFSA8A, SCMFSA6A, CAT_1, FUNCT_4, AMI_3, FUNCT_1, RELAT_1, CARD_1, AMI_5, UNIALG_2, SCMFSA7B, SCM_1, FUNCT_7, RELOC, ARYTM_1, SCMFSA6C, SCMFSA6B, SF_MASTR, SCMFSA_4, FUNCOP_1, SCMFSA8B, AMI_2, SCMFSA8C, CARD_3; notation TARSKI, XBOOLE_0, ENUMSET1, XCMPLX_0, XREAL_0, NAT_1, CQC_LANG, RELAT_1, FUNCT_1, FUNCT_4, FUNCT_7, CARD_1, STRUCT_0, AMI_1, AMI_3, AMI_5, SCMFSA_2, SCMFSA_4, SCMFSA_5, SCM_1, SF_MASTR, SCMFSA6A, SCMFSA6B, SCMFSA6C, SCMFSA7B, SCMFSA8A, SCMFSA8B, BINARITH; constructors ENUMSET1, AMI_5, SCM_1, SCMFSA_5, SCMFSA6A, SF_MASTR, SCMFSA6B, SCMFSA6C, SCMFSA8A, SCMFSA8B, BINARITH, MEMBERED; clusters FUNCT_1, PRELAMB, AMI_1, SCMFSA_2, SCMFSA_4, SCMFSA6A, SF_MASTR, SCMFSA6B, SCMFSA6C, SCMFSA7B, SCMFSA8A, INT_1, CQC_LANG, NAT_1, FRAENKEL, XREAL_0, MEMBERED; requirements REAL, NUMERALS, SUBSET, BOOLE, ARITHM; theorems FUNCT_4, SCMFSA_4, SCMFSA6A, AXIOMS, SCMFSA8B, AMI_1, SCMFSA6C, AMI_5, SCMFSA7B, SCMFSA8A, SCMFSA_7, SF_MASTR, NAT_1, SCMFSA6B, GRFUNC_1, SCMFSA_2, SCMFSA_5, FUNCT_1, FUNCT_7, REAL_1, AMI_3, TARSKI, GOBOARD9, REAL_2, CQC_THE1, SCM_1, ENUMSET1, CQC_LANG, RELAT_1, PRE_CIRC, SCMFSA_3, BINARITH, INT_1, SQUARE_1, XBOOLE_0, XBOOLE_1, XCMPLX_1; schemes NAT_1; begin :: Preliminaries reserve m for Nat; set A = the Instruction-Locations of SCM+FSA; set D = Int-Locations \/ FinSeq-Locations; definition cluster pseudo-paraclosed Macro-Instruction; existence proof consider K being pseudo-paraclosed Macro-Instruction; take K; thus thesis; end; end; canceled; theorem Th2: ::T4425 ** n.t for s being State of SCM+FSA,P being initial FinPartState of SCM+FSA st P is_pseudo-closed_on s for k being Nat st (for n being Nat st n <= k holds IC (Computation (s +* (P +* Start-At insloc 0))).n in dom P) holds k < pseudo-LifeSpan(s,P) proof let s be State of SCM+FSA; let P be initial FinPartState of SCM+FSA; assume A1: P is_pseudo-closed_on s; let k be Nat; assume A2: for n being Nat st n <= k holds IC (Computation (s +* (P +* Start-At insloc 0))).n in dom P; IC (Computation (s +* (P +* Start-At insloc 0))).pseudo-LifeSpan(s,P) = insloc card ProgramPart P by A1,SCMFSA8A:def 5; then not IC (Computation (s +* (P +* Start-At insloc 0))).pseudo-LifeSpan(s ,P) in dom ProgramPart P by SCMFSA6A:15; then A3: not IC (Computation (s +* (P +* Start-At insloc 0))).pseudo-LifeSpan(s ,P) in dom P by AMI_5:73; assume pseudo-LifeSpan(s,P) <= k; hence contradiction by A2,A3; end; canceled 3; theorem Th6: ::T210837 for f be Function, x be set st x in dom f holds f +* (x .--> f.x) = f proof let f be Function; let x be set; assume x in dom f; hence f +* (x .--> f.x) = f +*(x,f.x) by FUNCT_7:def 3 .= f by FUNCT_7:37; end; theorem Th7: ::TMP12 for l being Instruction-Location of SCM+FSA holds l + 0 = l proof let l be Instruction-Location of SCM+FSA; consider m being Nat such that A1: l = insloc m & l + 0 = insloc (m + 0) by SCMFSA_4:def 1; thus l + 0 = l by A1; end; theorem Th8: ::TMP19 for i being Instruction of SCM+FSA holds IncAddr(i,0) = i proof let i be Instruction of SCM+FSA; A1: InsCode i <= 11 + 1 by SCMFSA_2:35; A2: InsCode i <= 10+1 implies InsCode i <= 10 or InsCode i = 11 by NAT_1:26; A3: InsCode i <= 9+1 implies InsCode i <= 8+1 or InsCode i = 10 by NAT_1:26; A4: InsCode i <= 8+1 implies InsCode i <= 7+1 or InsCode i = 9 by NAT_1:26; per cases by A1,A2,A3,A4,CQC_THE1:9,NAT_1:26; suppose InsCode i = 0; then i = halt SCM+FSA by SCMFSA_2:122; hence IncAddr(i,0) = i by SCMFSA_4:8; suppose InsCode i = 1; then ex a,b being Int-Location st i = a:=b by SCMFSA_2:54; hence IncAddr(i,0) = i by SCMFSA_4:9; suppose InsCode i = 2; then ex a,b being Int-Location st i = AddTo(a,b) by SCMFSA_2:55; hence IncAddr(i,0) = i by SCMFSA_4:10; suppose InsCode i = 3; then ex a,b being Int-Location st i = SubFrom(a,b) by SCMFSA_2:56; hence IncAddr(i,0) = i by SCMFSA_4:11; suppose InsCode i = 4; then ex a,b being Int-Location st i = MultBy(a,b) by SCMFSA_2:57; hence IncAddr(i,0) = i by SCMFSA_4:12; suppose InsCode i = 5; then ex a,b being Int-Location st i = Divide(a,b) by SCMFSA_2:58; hence IncAddr(i,0) = i by SCMFSA_4:13; suppose InsCode i = 6; then consider l being Instruction-Location of SCM+FSA such that A5: i = goto l by SCMFSA_2:59; thus IncAddr(i,0) = goto (l + 0) by A5,SCMFSA_4:14 .= i by A5,Th7; suppose InsCode i = 7; then consider l being Instruction-Location of SCM+FSA,a being Int-Location such that A6: i = a =0_goto l by SCMFSA_2:60; thus IncAddr(i,0) = a =0_goto (l + 0) by A6,SCMFSA_4:15 .= i by A6,Th7; suppose InsCode i = 8; then consider l being Instruction-Location of SCM+FSA,a being Int-Location such that A7: i = a >0_goto l by SCMFSA_2:61; thus IncAddr(i,0) = a >0_goto (l + 0) by A7,SCMFSA_4:16 .= i by A7,Th7; suppose InsCode i = 9; then ex a,b being Int-Location,f being FinSeq-Location st i = b:=(f,a) by SCMFSA_2:62; hence IncAddr(i,0) = i by SCMFSA_4:17; suppose InsCode i = 10; then ex a,b being Int-Location,f being FinSeq-Location st i = (f,a):=b by SCMFSA_2:63; hence IncAddr(i,0) = i by SCMFSA_4:18; suppose InsCode i = 11; then ex a being Int-Location,f being FinSeq-Location st i = a:=len f by SCMFSA_2:64; hence IncAddr(i,0) = i by SCMFSA_4:19; suppose InsCode i = 12; then ex a being Int-Location,f being FinSeq-Location st i = f:=<0,...,0>a by SCMFSA_2:65; hence IncAddr(i,0) = i by SCMFSA_4:20; end; theorem Th9: ::TMP13 for P being programmed FinPartState of SCM+FSA holds ProgramPart Relocated(P,0) = P proof let P be programmed FinPartState of SCM+FSA; now let x be set; hereby assume A1: x in dom P; dom P c= the Instruction-Locations of SCM+FSA by AMI_3:def 13; then consider n being Nat such that A2: x = insloc n by A1,SCMFSA_2:21; thus x in {insloc m:insloc m in dom P} by A1,A2; end; assume x in {insloc m:insloc m in dom P}; then consider m being Nat such that A3: x = insloc m & insloc m in dom P; thus x in dom P by A3; end; then A4: dom P = {insloc m:insloc m in dom P} by TARSKI:2; A5: dom ProgramPart P = dom P by AMI_5:72; now let x be set; A6: dom ProgramPart Relocated(P,0) = {insloc (m + 0):insloc m in dom P} by A5,SCMFSA_5:3; hereby assume x in dom ProgramPart Relocated(P,0); then consider n be Nat such that A7: x = insloc (n + 0) & insloc n in dom P by A6; thus x in {insloc m:insloc m in dom P} by A7; end; assume x in {insloc m:insloc m in dom P}; then consider m being Nat such that A8: x = insloc m & insloc m in dom P; x = insloc (m + 0) & insloc m in dom P by A8; hence x in dom ProgramPart Relocated(P,0) by A6; end; then A9: dom ProgramPart Relocated(P,0) = {insloc m:insloc m in dom P} by TARSKI:2; now let x be set; assume x in {insloc m:insloc m in dom P}; then consider n being Nat such that A10: x = insloc n and A11: insloc n in dom P; A12: insloc n in dom ProgramPart P by A11,AMI_5:72; dom Shift(ProgramPart P,0) = {insloc (m + 0):insloc m in dom ProgramPart P} by SCMFSA_4:def 7; then A13: insloc (n + 0) in dom Shift(ProgramPart P,0) by A5,A11; then A14: pi(Shift(ProgramPart P,0),insloc (n + 0)) = Shift(ProgramPart P,0).insloc (n + 0) by AMI_5:def 5 .= Shift(ProgramPart P,0).(insloc n + 0) by SCMFSA_4:def 1 .= (ProgramPart P).insloc n by A12,SCMFSA_4:30 .= P.insloc n by AMI_5:72; then consider i being Instruction of SCM+FSA such that A15: i = P.insloc n; thus (ProgramPart Relocated(P,0)).x = IncAddr(Shift(ProgramPart P,0),0).insloc (n + 0) by A10,SCMFSA_5:2 .= IncAddr(i,0) by A13,A14,A15,SCMFSA_4:def 6 .= P.x by A10,A15,Th8; end; hence ProgramPart Relocated(P,0) = P by A4,A9,FUNCT_1:9; end; theorem Th10: ::TMP15' for P,Q being FinPartState of SCM+FSA st P c= Q holds ProgramPart P c= ProgramPart Q proof let P,Q be FinPartState of SCM+FSA; assume A1: P c= Q; then A2: dom P c= dom Q by GRFUNC_1:8; A3: ProgramPart P = P | A & ProgramPart Q = Q | A by AMI_5:def 6; then A4: dom ProgramPart P = dom P /\ A & dom ProgramPart Q = dom Q /\ A by RELAT_1:90; then A5: dom ProgramPart P c= dom ProgramPart Q by A2,XBOOLE_1:26; now let x be set; assume A6: x in dom ProgramPart P; then A7: x in dom P by A4,XBOOLE_0:def 3; thus (ProgramPart P).x = P.x by A3,A6,FUNCT_1:68 .= Q.x by A1,A7,GRFUNC_1:8 .= (ProgramPart Q).x by A3,A5,A6,FUNCT_1:68; end; hence ProgramPart P c= ProgramPart Q by A5,GRFUNC_1:8; end; theorem Th11: ::TMP18 for P,Q being programmed FinPartState of SCM+FSA, k being Nat st P c= Q holds Shift(P,k) c= Shift(Q,k) proof let P,Q be programmed FinPartState of SCM+FSA; let k be Nat; assume A1: P c= Q; then A2: dom P c= dom Q by GRFUNC_1:8; A3: dom Shift(P,k) = {insloc (m + k):insloc m in dom P} by SCMFSA_4:def 7; A4: dom Shift(Q,k) = {insloc (m + k):insloc m in dom Q} by SCMFSA_4:def 7; now let x be set; assume x in dom Shift(P,k); then consider m1 being Nat such that A5: x = insloc (m1 + k) & insloc m1 in dom P by A3; thus x in dom Shift(Q,k) by A2,A4,A5; end; then A6: dom Shift(P,k) c= dom Shift(Q,k) by TARSKI:def 3; now let x be set; assume x in dom Shift(P,k); then consider m1 being Nat such that A7: x = insloc (m1 + k) & insloc m1 in dom P by A3; thus Shift(P,k).x = Shift(P,k).(insloc m1 + k) by A7,SCMFSA_4:def 1 .= P.insloc m1 by A7,SCMFSA_4:30 .= Q.insloc m1 by A1,A7,GRFUNC_1:8 .= Shift(Q,k).(insloc m1 + k) by A2,A7,SCMFSA_4:30 .= Shift(Q,k).x by A7,SCMFSA_4:def 1; end; hence Shift(P,k) c= Shift(Q,k) by A6,GRFUNC_1:8; end; theorem Th12: ::TMP15 for P,Q being FinPartState of SCM+FSA, k being Nat st P c= Q holds ProgramPart Relocated(P,k) c= ProgramPart Relocated(Q,k) proof let P,Q be FinPartState of SCM+FSA; let k be Nat; assume A1: P c= Q; set rP = Relocated(P,k); set rQ = Relocated(Q,k); A2: dom ProgramPart rP = {insloc (m + k):insloc m in dom ProgramPart P} by SCMFSA_5:3; A3: ProgramPart rP = IncAddr(Shift(ProgramPart P,k),k) by SCMFSA_5:2; A4: ProgramPart rQ = IncAddr(Shift(ProgramPart Q,k),k) by SCMFSA_5:2; A5: ProgramPart P c= ProgramPart Q by A1,Th10; then A6: Shift(ProgramPart P,k) c= Shift(ProgramPart Q,k) by Th11; then A7: dom Shift(ProgramPart P,k) c= dom Shift(ProgramPart Q,k) by GRFUNC_1:8 ; A8: dom ProgramPart P c= dom ProgramPart Q by A5,GRFUNC_1:8; now let x be set; assume x in dom ProgramPart rP; then x in dom Shift(ProgramPart P,k) by A3,SCMFSA_4:def 6; then x in dom Shift(ProgramPart Q,k) by A7; hence x in dom ProgramPart rQ by A4,SCMFSA_4:def 6; end; then A9: dom ProgramPart rP c= dom ProgramPart rQ by TARSKI:def 3; A10: dom Shift(ProgramPart P,k) = {insloc (m + k):insloc m in dom ProgramPart P} by SCMFSA_4:def 7; A11: dom Shift(ProgramPart Q,k) = {insloc (m + k):insloc m in dom ProgramPart Q} by SCMFSA_4:def 7; now let x be set; assume x in dom ProgramPart rP; then consider m1 being Nat such that A12: x = insloc (m1 + k) & insloc m1 in dom ProgramPart P by A2; A13: insloc (m1 + k) in dom Shift(ProgramPart P,k) & insloc (m1 + k) in dom Shift(ProgramPart Q,k) by A8,A10,A11,A12; then A14: pi(Shift(ProgramPart P,k),insloc (m1 + k)) = Shift(ProgramPart P,k).insloc (m1 + k) by AMI_5:def 5 .= Shift(ProgramPart Q,k).insloc (m1 + k) by A6,A13,GRFUNC_1:8 .= pi(Shift(ProgramPart Q,k),insloc (m1 + k)) by A13,AMI_5:def 5; thus (ProgramPart rP).x = IncAddr(Shift(ProgramPart P,k),k).insloc (m1 + k) by A12,SCMFSA_5:2 .= IncAddr(pi(Shift(ProgramPart Q,k),insloc (m1 + k)),k) by A13,A14, SCMFSA_4:def 6 .= IncAddr(Shift(ProgramPart Q,k),k).insloc (m1 + k) by A13,SCMFSA_4:def 6 .= (ProgramPart rQ).x by A12,SCMFSA_5:2; end; hence ProgramPart rP c= ProgramPart rQ by A9,GRFUNC_1:8; end; theorem Th13: ::TMP16 for I,J being Macro-Instruction, k being Nat st card I <= k & k < card I + card J holds for i being Instruction of SCM+FSA st i = J.insloc (k -' card I) holds (I ';' J).insloc k = IncAddr(i,card I) proof let I,J be Macro-Instruction; let k be Nat; assume A1: card I <= k; assume k < card I + card J; then A2: k + 0 < card J + card I; let i be Instruction of SCM+FSA; assume A3: i = J.insloc (k -' card I); A4: I ';' J = Directed I +* ProgramPart Relocated(J,card I) by SCMFSA6A:def 4; k -' card I = k - card I by A1,SCMFSA_7:3; then A5: k -' card I < card J - 0 by A2,REAL_2:168; ProgramPart J = J by AMI_5:72; then A6: insloc (k -' card I) in dom ProgramPart J by A5,SCMFSA6A:15; A7: k -' card I + card I = k - card I + card I by A1,SCMFSA_7:3 .= k by XCMPLX_1:27; then A8: insloc (k -' card I) + card I = insloc k by SCMFSA_4:def 1; insloc k in {insloc (m + card I):insloc m in dom ProgramPart J} by A6,A7; then A9: insloc k in dom ProgramPart Relocated(J, card I) by SCMFSA_5:3; hence (I ';' J).insloc k = (ProgramPart Relocated(J,card I)).insloc k by A4,FUNCT_4:14 .= Relocated(J,card I).insloc k by A9,SCMFSA8A:49 .= IncAddr(i,card I) by A3,A6,A8,SCMFSA_5:7; end; theorem Th14: ::T22246 ---??? for s being State of SCM+FSA st s.intloc 0 = 1 & IC s = insloc 0 holds Initialize s = s proof let s be State of SCM+FSA; assume A1: s.intloc 0 = 1; assume IC s = insloc 0; then A2: s.IC SCM+FSA = insloc 0 by AMI_1:def 15; A3: intloc 0 in dom s & IC SCM+FSA in dom s by AMI_5:25,SCMFSA_2:66; thus Initialize s = s +* (intloc 0 .--> 1) +* Start-At insloc 0 by SCMFSA6C:def 3 .= s +* Start-At insloc 0 by A1,A3,Th6 .= s +* (IC SCM+FSA .--> insloc 0) by AMI_3:def 12 .= s by A2,A3,Th6; end; theorem Th15: ::T200922 for s being State of SCM+FSA holds Initialize Initialize s = Initialize s proof let s be State of SCM+FSA; (Initialize s).intloc 0 = 1 & IC Initialize s = insloc 0 by SCMFSA6C:3; hence Initialize Initialize s = Initialize s by Th14; end; theorem Th16: ::TMP6 for s being State of SCM+FSA, I being Macro-Instruction holds s +* (Initialized I +* Start-At insloc 0) = Initialize s +* (I +* Start-At insloc 0) proof let s be State of SCM+FSA; let I be Macro-Instruction; Start-At insloc 0 c= Initialized I by SCMFSA6B:4; hence s +* (Initialized I +* Start-At insloc 0) = s +* Initialized I by AMI_5:10 .= Initialize s +* (I +* Start-At insloc 0) by SCMFSA8A:13; end; theorem Th17: ::T200938 for s being State of SCM+FSA, I being Macro-Instruction holds IExec(I,s) = IExec(I,Initialize s) proof let s be State of SCM+FSA; let I be Macro-Instruction; dom s = D \/ {IC SCM+FSA} \/ A by SCMFSA6A:34; then A1: A c= dom s by XBOOLE_1:7; dom Initialize s = D \/ {IC SCM+FSA} \/ A by SCMFSA6A:34; then A2: A c= dom Initialize s by XBOOLE_1:7; for x be set st x in A holds s.x = (Initialize s).x by SCMFSA6C:3; then A3: s | A = (Initialize s) | A by A1,A2,SCMFSA6A:9; thus IExec(I,s) = Result (s +* Initialized I) +* s | A by SCMFSA6B:def 1 .= Result (Initialize s +* Initialized I) +* s | A by SCMFSA8A:8 .= IExec(I,Initialize s) by A3,SCMFSA6B:def 1; end; theorem Th18: ::T26401 --- ??? for s being State of SCM+FSA, I being Macro-Instruction st s.intloc 0 = 1 holds s +* (I +* Start-At insloc 0) = s +* Initialized I proof let s be State of SCM+FSA; let I be Macro-Instruction; assume A1: s.intloc 0 = 1; intloc 0 in dom s by SCMFSA_2:66; then A2: s = s +* (intloc 0 .--> 1) by A1,Th6; thus s +* Initialized I = Initialize s +* (I +* Start-At insloc 0) by SCMFSA8A:13 .= Initialize s +* I +* Start-At insloc 0 by FUNCT_4:15 .= Initialize s +* Start-At insloc 0 +* I by SCMFSA6B:14 .= s +* Start-At insloc 0 +* Start-At insloc 0 +* I by A2,SCMFSA6C:def 3 .= s +* (Start-At insloc 0 +* Start-At insloc 0) +* I by FUNCT_4:15 .= s +* I +* Start-At insloc 0 by SCMFSA6B:14 .= s +* (I +* Start-At insloc 0) by FUNCT_4:15; end; theorem Th19: ::T24634 --- ??? for I being Macro-Instruction holds I +* Start-At insloc 0 c= Initialized I proof let I be Macro-Instruction; I c= Initialized I & Start-At insloc 0 c= Initialized I by SCMFSA6A:26,SCMFSA6B:4; hence I +* Start-At insloc 0 c= Initialized I by SCMFSA6A:6; end; theorem Th20: ::TMP10 for l being Instruction-Location of SCM+FSA, I being Macro-Instruction holds l in dom I iff l in dom Initialized I proof let l be Instruction-Location of SCM+FSA; let I be Macro-Instruction; A1: (Initialized I) | the Instruction-Locations of SCM+FSA = I by SCMFSA6A:33; A2: dom ((Initialized I) | the Instruction-Locations of SCM+FSA) c= dom Initialized I by FUNCT_1:76; A3: dom ((Initialized I) | the Instruction-Locations of SCM+FSA) = dom Initialized I /\ the Instruction-Locations of SCM+FSA by RELAT_1:90; thus l in dom I implies l in dom Initialized I by A1,A2; assume l in dom Initialized I; hence l in dom I by A1,A3,XBOOLE_0:def 3; end; theorem for s being State of SCM+FSA, I being Macro-Instruction holds Initialized I is_closed_on s iff I is_closed_on Initialize s proof let s be State of SCM+FSA; let I be Macro-Instruction; hereby assume A1: Initialized I is_closed_on s; now let k be Nat; IC (Computation (s +* (Initialized I +* Start-At insloc 0))).k in dom Initialized I by A1,SCMFSA7B:def 7; then IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k in dom Initialized I by Th16; hence IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k in dom I by Th20; end; hence I is_closed_on Initialize s by SCMFSA7B:def 7; end; assume A2: I is_closed_on Initialize s; now let k be Nat; IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k in dom I by A2,SCMFSA7B:def 7; then IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k in dom Initialized I by Th20; hence IC (Computation (s +* (Initialized I +* Start-At insloc 0))).k in dom Initialized I by Th16; end; hence Initialized I is_closed_on s by SCMFSA7B:def 7; end; theorem Th22: ::TMP5 for s being State of SCM+FSA, I being Macro-Instruction holds Initialized I is_halting_on s iff I is_halting_on Initialize s proof let s be State of SCM+FSA; let I be Macro-Instruction; hereby assume Initialized I is_halting_on s; then s +* (Initialized I +* Start-At insloc 0) is halting by SCMFSA7B:def 8; then Initialize s +* (I +* Start-At insloc 0) is halting by Th16; hence I is_halting_on Initialize s by SCMFSA7B:def 8; end; assume I is_halting_on Initialize s; then Initialize s +* (I +* Start-At insloc 0) is halting by SCMFSA7B:def 8; then s +* (Initialized I +* Start-At insloc 0) is halting by Th16; hence Initialized I is_halting_on s by SCMFSA7B:def 8; end; theorem for I being Macro-Instruction holds (for s being State of SCM+FSA holds I is_halting_on Initialize s) implies Initialized I is halting proof let I be Macro-Instruction; assume A1: for s being State of SCM+FSA holds I is_halting_on Initialize s; now let s be State of SCM+FSA; assume A2: Initialized I c= s; I is_halting_on Initialize s by A1; then Initialize s +* (I +* Start-At insloc 0) is halting by SCMFSA7B:def 8; then s +* Initialized I is halting by SCMFSA8A:13; hence s is halting by A2,AMI_5:10; end; hence Initialized I is halting by AMI_1:def 26; end; theorem Th24: ::TMP3' for I being Macro-Instruction holds (for s being State of SCM+FSA holds Initialized I is_halting_on s) implies Initialized I is halting proof let I be Macro-Instruction; assume A1: for s being State of SCM+FSA holds Initialized I is_halting_on s; now let s be State of SCM+FSA; assume A2: Initialized I c= s; Initialized I is_halting_on s by A1; then A3: s +* (Initialized I +* Start-At insloc 0) is halting by SCMFSA7B: def 8; Start-At insloc 0 c= Initialized I by SCMFSA6B:4; then s +* (Initialized I +* Start-At insloc 0) = s +* Initialized I by AMI_5:10; hence s is halting by A2,A3,AMI_5:10; end; hence Initialized I is halting by AMI_1:def 26; end; theorem Th25: ::TMP9 for I being Macro-Instruction holds ProgramPart Initialized I = I proof let I be Macro-Instruction; thus ProgramPart Initialized I = (Initialized I) | the Instruction-Locations of SCM+FSA by AMI_5:def 6 .= I by SCMFSA6A:33; end; theorem Th26: ::T18750 for s being State of SCM+FSA, I being Macro-Instruction, l being Instruction-Location of SCM+FSA, x being set holds x in dom I implies I.x = (s +* (I +* Start-At l)).x proof let s be State of SCM+FSA; let I be Macro-Instruction; let l be Instruction-Location of SCM+FSA; let x be set; assume A1: x in dom I; dom I c= the Instruction-Locations of SCM+FSA by AMI_3:def 13; then reconsider y = x as Instruction-Location of SCM+FSA by A1; A2: not y in dom Start-At l by SCMFSA6B:11; x in dom (I +* Start-At l) by A1,FUNCT_4:13; hence (s +* (I +* Start-At l)).x = (I +* Start-At l).x by FUNCT_4:14 .= I.x by A2,FUNCT_4:12; end; theorem Th27: ::T22246' ---??? for s being State of SCM+FSA st s.intloc 0 = 1 holds (Initialize s) | (Int-Locations \/ FinSeq-Locations) = s | (Int-Locations \/ FinSeq-Locations) proof let s be State of SCM+FSA; assume A1: s.intloc 0 = 1; A2: intloc 0 in dom s & IC SCM+FSA in dom s by AMI_5:25,SCMFSA_2:66; Initialize s = s +* (intloc 0 .--> 1) +* Start-At insloc 0 by SCMFSA6C:def 3 .= s +* Start-At insloc 0 by A1,A2,Th6; hence thesis by SCMFSA8A:10; end; theorem Th28: ::T210615 for s being State of SCM+FSA, I being Macro-Instruction, a being Int-Location, l being Instruction-Location of SCM+FSA holds (s +* (I +* Start-At l)).a = s.a proof let s be State of SCM+FSA; let I be Macro-Instruction; let a be Int-Location; let l be Instruction-Location of SCM+FSA; a in dom s & not a in dom (I +* Start-At l) by SCMFSA6B:12,SCMFSA_2:66; hence (s +* (I +* Start-At l)).a = s.a by FUNCT_4:12; end; theorem for I being programmed FinPartState of SCM+FSA, l being Instruction-Location of SCM+FSA holds IC SCM+FSA in dom (I +* Start-At l) proof let I be programmed FinPartState of SCM+FSA; let l be Instruction-Location of SCM+FSA; dom Start-At l = {IC SCM+FSA} by AMI_3:34; then IC SCM+FSA in dom Start-At l by TARSKI:def 1; hence IC SCM+FSA in dom (I +* Start-At l) by FUNCT_4:13; end; theorem for I being programmed FinPartState of SCM+FSA, l being Instruction-Location of SCM+FSA holds (I +* Start-At l).IC SCM+FSA = l proof let I be programmed FinPartState of SCM+FSA; let l be Instruction-Location of SCM+FSA; dom Start-At l = {IC SCM+FSA} by AMI_3:34; then A1: IC SCM+FSA in dom Start-At l by TARSKI:def 1; Start-At l = IC SCM+FSA .--> l by AMI_3:def 12; then (Start-At l).IC SCM+FSA = l by CQC_LANG:6; hence (I +* Start-At l).IC SCM+FSA = l by A1,FUNCT_4:14; end; theorem Th31: ::T4633 -- ?? for s being State of SCM+FSA, P being FinPartState of SCM+FSA, l being Instruction-Location of SCM+FSA holds IC (s +* (P +* Start-At l)) = l proof let s be State of SCM+FSA; let I be FinPartState of SCM+FSA; let l be Instruction-Location of SCM+FSA; thus IC (s +* (I +* Start-At l)) = IC (s +* I +* Start-At l) by FUNCT_4:15 .= l by AMI_5:79; end; theorem Th32: ::T30408 for s being State of SCM+FSA, i being Instruction of SCM+FSA st InsCode i in {0,6,7,8} holds Exec(i,s) | (Int-Locations \/ FinSeq-Locations) = s | (Int-Locations \/ FinSeq-Locations) proof let s be State of SCM+FSA; let i be Instruction of SCM+FSA; assume A1: InsCode i in {0,6,7,8}; now let a be Int-Location; let f be FinSeq-Location; per cases by A1,ENUMSET1:18; suppose InsCode i = 0; then i = halt SCM+FSA by SCMFSA_2:122; hence Exec(i,s).a = s.a & Exec(i,s).f = s.f by AMI_1:def 8; suppose InsCode i = 6; then consider lb being Instruction-Location of SCM+FSA such that A2: i = goto lb by SCMFSA_2:59; thus Exec(i,s).a = s.a & Exec(i,s).f = s.f by A2,SCMFSA_2:95; suppose InsCode i = 7; then consider lb being Instruction-Location of SCM+FSA, b being Int-Location such that A3: i = b=0_goto lb by SCMFSA_2:60; thus Exec(i,s).a = s.a & Exec(i,s).f = s.f by A3,SCMFSA_2:96; suppose InsCode i = 8; then consider lb being Instruction-Location of SCM+FSA, b being Int-Location such that A4: i = b>0_goto lb by SCMFSA_2:61; thus Exec(i,s).a = s.a & Exec(i,s).f = s.f by A4,SCMFSA_2:97; end; hence thesis by SCMFSA6A:38; end; theorem Th33: ::T271245 --- ??? for s1,s2 being State of SCM+FSA st s1.intloc 0 = s2.intloc 0 & ((for a being read-write Int-Location holds s1.a = s2.a) & for f being FinSeq-Location holds s1.f = s2.f) holds s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) proof let s1,s2 be State of SCM+FSA; assume A1: s1.intloc 0 = s2.intloc 0; assume A2: for a being read-write Int-Location holds s1.a = s2.a; assume A3: for f being FinSeq-Location holds s1.f = s2.f; A4: dom (s1 | D) = dom s1 /\ D by RELAT_1:90 .= (Int-Locations \/ FinSeq-Locations \/ {IC SCM+FSA} \/ the Instruction-Locations of SCM+FSA) /\ D by SCMFSA6A:34 .= dom s2 /\ D by SCMFSA6A:34 .= dom (s2 | D) by RELAT_1:90; now let x be set; assume A5: x in dom (s1 | D); then A6: x in dom s1 /\ D by RELAT_1:90; then A7: x in dom s1 & x in D by XBOOLE_0:def 3; per cases by A7,SCMFSA6A:35; suppose A8: x is Int-Location; hereby per cases; suppose A9: x is read-write Int-Location; thus (s1 | D).x = s1.x by A5,FUNCT_1:70 .= s2.x by A2,A9 .= (s2 | D).x by A4,A5,FUNCT_1:70; suppose A10: not x is read-write Int-Location; reconsider a = x as Int-Location by A8; a = intloc 0 by A10,SF_MASTR:def 5; hence (s1 | D).x = s2.a by A1,A5,FUNCT_1:70 .= (s2 | D).x by A4,A5,FUNCT_1:70; end; suppose A11: x is FinSeq-Location; thus (s1 | D).x = s1.x by A5,FUNCT_1:70 .= s2.x by A3,A11 .= (s2 | D).x by A4,A5,FUNCT_1:70; suppose A12: x = IC SCM+FSA; assume not (s1 | D).x = (s2 | D).x; thus contradiction by A6,A12,SCMFSA6A:37,XBOOLE_0:def 3; suppose A13: x is Instruction-Location of SCM+FSA; assume not (s1 | D).x = (s2 | D).x; thus contradiction by A7,A13,SCMFSA6A:37; end; then s1 | D c= s2 | D by A4,GRFUNC_1:8; hence s1 | D = s2 | D by A4,GRFUNC_1:9; end; theorem for s being State of SCM+FSA,P being programmed FinPartState of SCM+FSA holds (s +* P) | (Int-Locations \/ FinSeq-Locations) = s | (Int-Locations \/ FinSeq-Locations) proof let s be State of SCM+FSA; let P be programmed FinPartState of SCM+FSA; A1: dom P c= A by AMI_3:def 13; A misses D by SCMFSA_2:13,14,XBOOLE_1:70; then dom P misses D by A1,XBOOLE_1:63; hence thesis by AMI_5:7; end; theorem Th35: ::T1643 --- ?? for s,ss being State of SCM+FSA holds (s +* ss | the Instruction-Locations of SCM+FSA) | (Int-Locations \/ FinSeq-Locations) = s | (Int-Locations \/ FinSeq-Locations) proof let s,ss be State of SCM+FSA; dom (ss | A) = dom ss /\ A by RELAT_1:90 .= (Int-Locations \/ FinSeq-Locations \/ {IC SCM+FSA} \/ A) /\ A by SCMFSA6A:34 .= A by XBOOLE_1:21; then dom (ss | A) misses D by SCMFSA_2:13,14,XBOOLE_1:70; hence thesis by AMI_5:7; end; theorem Th36: ::T27643 --- ??? for s being State of SCM+FSA holds (Initialize s) | the Instruction-Locations of SCM+FSA = s | the Instruction-Locations of SCM+FSA proof let s be State of SCM+FSA; A1: Initialize s = s +* (intloc 0 .--> 1) +* Start-At insloc 0 by SCMFSA6C:def 3 .= s +* ((intloc 0 .--> 1) +* Start-At insloc 0) by FUNCT_4:15; now let x be set; assume A2: x in A; then A3: not x in dom Start-At insloc 0 by SCMFSA6B:11; x <> intloc 0 & dom (intloc 0 .--> 1) = {intloc 0} by A2,CQC_LANG:5,SCMFSA_2:84; then not x in dom (intloc 0 .--> 1) by TARSKI:def 1; hence not x in dom (intloc 0 .--> 1) \/ dom Start-At insloc 0 by A3,XBOOLE_0:def 2; end; then dom (intloc 0 .--> 1) \/ dom Start-At insloc 0 misses A by XBOOLE_0:3; then dom ((intloc 0 .--> 1) +* Start-At insloc 0) misses A by FUNCT_4:def 1 ; hence thesis by A1,SCMFSA8A:2; end; theorem Th37: ::T26530 --- ??? for s,ss being State of SCM+FSA, I being Macro-Instruction holds (ss +* (s | (the Instruction-Locations of SCM+FSA))) | (Int-Locations \/ FinSeq-Locations) = ss | (Int-Locations \/ FinSeq-Locations) proof let s,ss be State of SCM+FSA; let I be Macro-Instruction; dom (s | A) = A & A misses D by SCMFSA8A:3,SCMFSA_2:13,14,XBOOLE_1:70 ; hence thesis by SCMFSA8A:2; end; theorem Th38: ::T27500 ue for s being State of SCM+FSA holds IExec(SCM+FSA-Stop,s) = Initialize s +* Start-At insloc 0 proof let s be State of SCM+FSA; set s1 = Initialize s +* (SCM+FSA-Stop +* Start-At insloc 0); A1: s +* Initialized SCM+FSA-Stop = s1 by SCMFSA8A:13; then Initialized SCM+FSA-Stop c= s1 by FUNCT_4:26; then IC s1 = insloc 0 by SCMFSA6B:34; then CurInstr s1 = s1.insloc 0 by AMI_1:def 17 .= SCM+FSA-Stop.insloc 0 by Th26,SCMFSA8A:16; then A2: CurInstr s1 = halt SCM+FSA & s1 = (Computation s1).0 by AMI_1:def 19,SCMFSA8A:16; then A3: s1 is halting by AMI_1:def 20; A4: IExec(SCM+FSA-Stop,s) = Result (s +* Initialized SCM+FSA-Stop) +* s | A by SCMFSA6B:def 1; then A5: IExec(SCM+FSA-Stop,s) = s1 +* s | A by A1,A2,A3,AMI_1:def 22; then A6: IExec(SCM+FSA-Stop,s) | D = s1 | D by Th37 .= (Initialize s) | D by SCMFSA8A:11; hereby A7: dom IExec(SCM+FSA-Stop,s) = the carrier of SCM+FSA by AMI_3:36 .= dom (Initialize s +* Start-At insloc 0) by AMI_3:36; A8: dom Start-At insloc 0 = {IC SCM+FSA} by AMI_3:34; now let x be set; assume A9: x in dom IExec(SCM+FSA-Stop,s); per cases by A9,SCMFSA6A:35; suppose A10: x is Int-Location; then A11: IExec(SCM+FSA-Stop,s).x = (Initialize s).x by A6,SCMFSA6A:38; x <> IC SCM+FSA by A10,SCMFSA_2:81; then not x in dom Start-At insloc 0 by A8,TARSKI:def 1; hence IExec(SCM+FSA-Stop,s).x = (Initialize s +* Start-At insloc 0).x by A11,FUNCT_4:12; suppose A12: x is FinSeq-Location; then A13: IExec(SCM+FSA-Stop,s).x = (Initialize s).x by A6,SCMFSA6A:38; x <> IC SCM+FSA by A12,SCMFSA_2:82; then not x in dom Start-At insloc 0 by A8,TARSKI:def 1; hence IExec(SCM+FSA-Stop,s).x = (Initialize s +* Start-At insloc 0).x by A13,FUNCT_4:12; suppose A14: x = IC SCM+FSA; then x in {IC SCM+FSA} by TARSKI:def 1; then A15: x in dom Start-At insloc 0 by AMI_3:34; not x in A by A14,AMI_1:48; then not x in dom s /\ A by XBOOLE_0:def 3; then x in dom s1 & not x in dom (s | A) by A14,RELAT_1:90,SCMFSA8B:1; hence IExec(SCM+FSA-Stop,s).x = s1.IC SCM+FSA by A5,A14,FUNCT_4:12 .= (Initialize s +* SCM+FSA-Stop +* Start-At insloc 0).IC SCM+FSA by FUNCT_4:15 .= (Start-At insloc 0).IC SCM+FSA by A14,A15,FUNCT_4:14 .= (Initialize s +* Start-At insloc 0).x by A14,A15,FUNCT_4:14; suppose A16: x is Instruction-Location of SCM+FSA; IExec(SCM+FSA-Stop,s) | A = s | A by A4,SCMFSA6A:40 .= (Initialize s) | A by Th36; then A17: IExec(SCM+FSA-Stop,s).x = (Initialize s).x by A16,SCMFSA6A:36; x <> IC SCM+FSA by A16,AMI_1:48; then not x in dom Start-At insloc 0 by A8,TARSKI:def 1; hence IExec(SCM+FSA-Stop,s).x = (Initialize s +* Start-At insloc 0).x by A17,FUNCT_4:12; end; hence thesis by A7,FUNCT_1:9; end; end; theorem Th39: ::T9x for s being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s holds insloc 0 in dom I proof let s be State of SCM+FSA; let I be Macro-Instruction; assume A1: I is_closed_on s; consider n being Nat such that A2: insloc n = IC (Computation (s +* (I +* Start-At insloc 0))).0 by SCMFSA_2:21; A3: insloc n in dom I by A1,A2,SCMFSA7B:def 7; per cases by NAT_1:19; suppose n = 0; hence insloc 0 in dom I by A1,A2,SCMFSA7B:def 7; suppose 0 < n; hence insloc 0 in dom I by A3,SCMFSA_4:def 4; end; theorem for s being State of SCM+FSA,I being paraclosed Macro-Instruction holds insloc 0 in dom I proof let s be State of SCM+FSA; let I be paraclosed Macro-Instruction; I is_closed_on s by SCMFSA7B:24; hence insloc 0 in dom I by Th39; end; theorem Th41: ::Tm1(@BBB8) for i being Instruction of SCM+FSA holds rng Macro i = {i,halt SCM+FSA} proof let i be Instruction of SCM+FSA; A1: Macro i = (insloc 0,insloc 1) --> (i,halt SCM+FSA) by SCMFSA6A:def 2; insloc 0 <> insloc 1 by SCMFSA_2:18; hence rng Macro i = {i,halt SCM+FSA} by A1,FUNCT_4:67; end; theorem Th42: ::T0x for s1,s2 being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s1 & I +* Start-At insloc 0 c= s1 for n being Nat st ProgramPart Relocated(I,n) c= s2 & IC s2 = insloc n & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) for i being Nat holds IC (Computation s1).i + n = IC (Computation s2).i & IncAddr(CurInstr ((Computation s1).i),n) = CurInstr ((Computation s2).i) & (Computation s1).i | (Int-Locations \/ FinSeq-Locations) = (Computation s2).i | (Int-Locations \/ FinSeq-Locations) proof let s1,s2 be State of SCM+FSA; let J be Macro-Instruction; assume A1: J is_closed_on s1; assume A2: J +* Start-At insloc 0 c= s1; set JAt = J +* Start-At insloc 0; let n be Nat; assume that A3: ProgramPart Relocated(J,n) c= s2 and A4: IC s2 = insloc n and A5: s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations); set C1 = Computation s1; set C2 = Computation s2; let i be Nat; defpred P[Nat] means IC C1.$1 + n = IC C2.$1 & IncAddr(CurInstr (C1.$1),n) = CurInstr (C2.$1) & C1.$1 | (Int-Locations \/ FinSeq-Locations) = C2.$1 | (Int-Locations \/ FinSeq-Locations); A6: ProgramPart Relocated(J,n) = Relocated(J,n) | the Instruction-Locations of SCM+FSA by AMI_5:def 6; A7: P[0] proof A8: IC SCM+FSA in dom JAt by SF_MASTR:65; insloc 0 in dom J by A1,Th39; then insloc 0 + n in dom Relocated(J,n) by SCMFSA_5:4; then insloc 0 + n in dom ProgramPart Relocated(J,n) by AMI_5:73; then A9: insloc (0 + n) in dom ProgramPart Relocated(J,n) by SCMFSA_4:def 1; IC C1.0 = IC s1 by AMI_1:def 19 .= s1.IC SCM+FSA by AMI_1:def 15 .= (JAt).IC SCM+FSA by A2,A8,GRFUNC_1:8 .= insloc 0 by SF_MASTR:66; hence IC C1.0 + n = insloc (0 + n) by SCMFSA_4:def 1 .= IC C2.0 by A4,AMI_1:def 19; dom J misses dom Start-At insloc 0 by SF_MASTR:64; then A10: J c= JAt by FUNCT_4:33; then A11: dom J c= dom JAt by GRFUNC_1:8; A12: insloc 0 in dom J by A1,Th39; A13: s1.IC s1 = s1.(s1.IC SCM+FSA) by AMI_1:def 15 .= s1.((JAt).IC SCM+FSA) by A2,A8,GRFUNC_1:8 .= s1.insloc 0 by SF_MASTR:66 .= (JAt).insloc 0 by A2,A11,A12,GRFUNC_1:8 .= J.insloc 0 by A10,A12,GRFUNC_1:8; ProgramPart J = J by AMI_5:72; then A14: insloc 0 in dom ProgramPart J by A1,Th39; thus IncAddr(CurInstr (C1.0),n) = IncAddr(CurInstr s1,n) by AMI_1:def 19 .= IncAddr(s1.IC s1,n) by AMI_1:def 17 .= Relocated(J,n).(insloc 0 + n) by A13,A14,SCMFSA_5:7 .= Relocated(J,n).insloc (0 + n) by SCMFSA_4:def 1 .= (ProgramPart Relocated(J,n)).insloc n by A6,FUNCT_1:72 .= s2.IC s2 by A3,A4,A9,GRFUNC_1:8 .= CurInstr s2 by AMI_1:def 17 .= CurInstr (C2.0) by AMI_1:def 19; thus C1.0 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) by A5,AMI_1:def 19 .= C2.0 | (Int-Locations \/ FinSeq-Locations) by AMI_1:def 19; end; A15: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume A16: P[k]; A17: C1.(k + 1) = Following C1.k by AMI_1:def 19 .= Exec(CurInstr C1.k,C1.k) by AMI_1:def 18; A18: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18; hence A19: IC C1.(k + 1) + n = IC C2.(k + 1) by A16,A17,SCMFSA6A:41; reconsider j = CurInstr C1.(k + 1) as Instruction of SCM+FSA; reconsider l = IC C1.(k + 1) as Instruction-Location of SCM+FSA; dom J misses dom Start-At insloc 0 by SF_MASTR:64; then A20: J c= JAt by FUNCT_4:33; then A21: dom J c= dom JAt by GRFUNC_1:8; s1 +* (J +* Start-At insloc 0) = s1 by A2,AMI_5:10; then A22: IC C1.(k + 1) in dom J by A1,SCMFSA7B:def 7; ProgramPart J = J | the Instruction-Locations of SCM+FSA by AMI_5:def 6; then dom ProgramPart J = dom J /\ the Instruction-Locations of SCM+FSA by FUNCT_1:68; then A23: l in dom ProgramPart J by A22,XBOOLE_0:def 3; A24: j = C1.(k + 1).IC C1.(k + 1) by AMI_1:def 17 .= s1.IC C1.(k + 1) by AMI_1:54 .= (JAt).IC C1.(k + 1) by A2,A21,A22,GRFUNC_1:8 .= J.l by A20,A22,GRFUNC_1:8; IC C2.(k + 1) in dom Relocated(J,n) by A19,A22,SCMFSA_5:4; then IC C2.(k + 1) in dom Relocated(J,n) /\ the Instruction-Locations of SCM+FSA by XBOOLE_0:def 3; then A25: IC C2.(k + 1) in dom ProgramPart Relocated(J,n) by A6,FUNCT_1:68; thus IncAddr(CurInstr C1.(k + 1),n) = Relocated(J,n).(l + n) by A23,A24,SCMFSA_5:7 .= (ProgramPart Relocated(J,n)).(IC C2.(k + 1)) by A6,A19,FUNCT_1:72 .= s2.IC C2.(k + 1) by A3,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) | (Int-Locations \/ FinSeq-Locations) = C2.(k + 1) | (Int-Locations \/ FinSeq-Locations) by A16,A17,A18,SCMFSA6A:41; end; for k being Nat holds P[k] from Ind(A7,A15); hence thesis; end; theorem Th43: ::T24441 for s1,s2 being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s1 & I +* Start-At insloc 0 c= s1 & I +* Start-At insloc 0 c= s2 & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) for i being Nat holds IC (Computation s1).i = IC (Computation s2).i & CurInstr (Computation s1).i = CurInstr (Computation s2).i & (Computation s1).i | (Int-Locations \/ FinSeq-Locations) = (Computation s2).i | (Int-Locations \/ FinSeq-Locations) proof let s1,s2 be State of SCM+FSA; let J be Macro-Instruction; assume A1: J is_closed_on s1 & J +* Start-At insloc 0 c= s1 & J +* Start-At insloc 0 c= s2 & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations); let i be Nat; A2: IC (Computation s1).i + 0 = IC (Computation s1).i by Th7; A3: s2 = s2 +* (J +* Start-At insloc 0) by A1,AMI_5:10 .= s2 +* J +* Start-At insloc 0 by FUNCT_4:15 .= s2 +* Start-At insloc 0 +* J by SCMFSA6B:14; ProgramPart Relocated(J,0) = J by Th9; then A4: ProgramPart Relocated(J,0) c= s2 by A3,FUNCT_4:26; A5: IncAddr(CurInstr (Computation s1).i,0) = CurInstr (Computation s1).i by Th8; IC s2 = IC (s2 +* (J +* Start-At insloc 0)) by A1,AMI_5:10 .= IC (s2 +* J +* Start-At insloc 0) by FUNCT_4:15 .= insloc 0 by AMI_5:79; hence IC (Computation s1).i = IC (Computation s2).i & CurInstr (Computation s1).i = CurInstr (Computation s2).i & (Computation s1).i | (Int-Locations \/ FinSeq-Locations) = (Computation s2).i | (Int-Locations \/ FinSeq-Locations) by A1,A2,A4,A5,Th42; end; theorem Th44: ::T24441' for s1,s2 being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s1 & I is_halting_on s1 & I +* Start-At insloc 0 c= s1 & I +* Start-At insloc 0 c= s2 & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) holds LifeSpan s1 = LifeSpan s2 proof let s1,s2 be State of SCM+FSA; let J be Macro-Instruction; assume A1: J is_closed_on s1 & J is_halting_on s1 & J +* Start-At insloc 0 c= s1 & J +* Start-At insloc 0 c= s2 & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations); then s1 = s1 +* (J +* Start-At insloc 0) by AMI_5:10; then A2: s1 is halting by A1,SCMFSA7B:def 8; then CurInstr (Computation s1).(LifeSpan s1) = halt SCM+FSA & for k being Nat st CurInstr (Computation s1).k = halt SCM+FSA holds LifeSpan s1 <= k by SCM_1:def 2; then A3: CurInstr (Computation s2).(LifeSpan s1) = halt SCM+FSA by A1,Th43; then A4: s2 is halting by AMI_1:def 20; now let k be Nat; assume CurInstr (Computation s2).k = halt SCM+FSA; then CurInstr (Computation s1).k = halt SCM+FSA by A1,Th43; hence LifeSpan s1 <= k by A2,SCM_1:def 2; end; hence LifeSpan s1 = LifeSpan s2 by A3,A4,SCM_1:def 2; end; theorem Th45: ::T27835 for s1,s2 being State of SCM+FSA,I being Macro-Instruction st s1.intloc 0 = 1 & I is_closed_on s1 & I is_halting_on s1 & ((for a being read-write Int-Location holds s1.a = s2.a) & for f being FinSeq-Location holds s1.f = s2.f) holds IExec(I,s1) | (Int-Locations \/ FinSeq-Locations) = IExec(I,s2) | (Int-Locations \/ FinSeq-Locations) proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; set s11 = s1 +* Initialized I; set s21 = s2 +* Initialized I; assume A1: s1.intloc 0 = 1; assume A2: I is_closed_on s1 & I is_halting_on s1; assume A3: for a being read-write Int-Location holds s1.a = s2.a; assume A4: for f being FinSeq-Location holds s1.f = s2.f; A5: Initialized I c= s11 & Initialized I c= s21 by FUNCT_4:26; A6: s11 = s1 +* (I +* Start-At insloc 0) by A1,Th18; then s11 | D = s1 | D by SCMFSA8A:11; then A7: I is_closed_on s11 & I is_halting_on s11 by A2,SCMFSA8B:8; A8: now let a be read-write Int-Location; A9: a in dom s1 & a in dom s2 & not a in dom Initialized I by SCMFSA6A:48,SCMFSA_2:66; hence s11.a = s1.a by FUNCT_4:12 .= s2.a by A3 .= s21.a by A9,FUNCT_4:12; end; A10: now let f be FinSeq-Location; A11: f in dom s1 & f in dom s2 & not f in dom Initialized I by SCMFSA6A:49,SCMFSA_2:67; hence s11.f = s1.f by FUNCT_4:12 .= s2.f by A4 .= s21.f by A11,FUNCT_4:12; end; A12: intloc 0 in dom Initialized I by SCMFSA6A:45; then s11.intloc 0 = (Initialized I).intloc 0 by FUNCT_4:14 .= s21.intloc 0 by A12,FUNCT_4:14; then A13: s11 | D = s21 | D by A8,A10,Th33; I +* Start-At insloc 0 c= Initialized I by Th19; then A14: I +* Start-At insloc 0 c= s11 & I +* Start-At insloc 0 c= s21 by A5,XBOOLE_1:1; then A15: LifeSpan s11 = LifeSpan s21 by A7,A13,Th44; A16: s11 is halting by A2,A6,SCMFSA7B:def 8; then CurInstr (Computation s11).(LifeSpan s11) = halt SCM+FSA & for k being Nat st CurInstr (Computation s11).k = halt SCM+FSA holds LifeSpan s11 <= k by SCM_1:def 2; then CurInstr (Computation s21).(LifeSpan s11) = halt SCM+FSA by A7,A13,A14,Th43; then A17: s21 is halting by AMI_1:def 20; thus IExec(I,s1) | (Int-Locations \/ FinSeq-Locations) = (Result s11) | D by SCMFSA8B:35 .= (Computation s11).(LifeSpan s11) | D by A16,SCMFSA6B:16 .= (Computation s21).(LifeSpan s11) | D by A7,A13,A14,Th43 .= (Result s21) | D by A15,A17,SCMFSA6B:16 .= IExec(I,s2) | (Int-Locations \/ FinSeq-Locations) by SCMFSA8B:35; end; theorem Th46: ::T27835' for s1,s2 being State of SCM+FSA,I being Macro-Instruction st s1.intloc 0 = 1 & I is_closed_on s1 & I is_halting_on s1 & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) holds IExec(I,s1) | (Int-Locations \/ FinSeq-Locations) = IExec(I,s2) | (Int-Locations \/ FinSeq-Locations) proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; set s11 = s1 +* Initialized I; set s21 = s2 +* Initialized I; assume A1: s1.intloc 0 = 1 & I is_closed_on s1 & I is_halting_on s1 & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations); A2: Initialized I c= s11 & Initialized I c= s21 by FUNCT_4:26; s2.intloc 0 = 1 by A1,SCMFSA6A:38; then A3: s11 = s1 +* (I +* Start-At insloc 0) & s21 = s2 +* (I +* Start-At insloc 0) by A1,Th18; then A4: s11 | D = s1 | D by SCMFSA8A:11; then A5: I is_closed_on s11 & I is_halting_on s11 by A1,SCMFSA8B:8; A6: s11 | D = s21 | D by A1,A3,A4,SCMFSA8A:11; I +* Start-At insloc 0 c= Initialized I by Th19; then A7: I +* Start-At insloc 0 c= s11 & I +* Start-At insloc 0 c= s21 by A2,XBOOLE_1:1; then A8: LifeSpan s11 = LifeSpan s21 by A5,A6,Th44; A9: s11 is halting by A1,A3,SCMFSA7B:def 8; then CurInstr (Computation s11).(LifeSpan s11) = halt SCM+FSA & for k being Nat st CurInstr (Computation s11).k = halt SCM+FSA holds LifeSpan s11 <= k by SCM_1:def 2; then CurInstr (Computation s21).(LifeSpan s11) = halt SCM+FSA by A5,A6,A7,Th43; then A10: s21 is halting by AMI_1:def 20; thus IExec(I,s1) | (Int-Locations \/ FinSeq-Locations) = (Result s11) | D by SCMFSA8B:35 .= (Computation s11).(LifeSpan s11) | D by A9,SCMFSA6B:16 .= (Computation s21).(LifeSpan s11) | D by A5,A6,A7,Th43 .= (Result s21) | D by A8,A10,SCMFSA6B:16 .= IExec(I,s2) | (Int-Locations \/ FinSeq-Locations) by SCMFSA8B:35; end; definition let I be Macro-Instruction; cluster Initialized I -> initial; correctness proof now let m,n be Nat; assume insloc n in dom Initialized I; then A1: insloc n in dom I by Th20; I c= Initialized I by SCMFSA6A:26; then A2: dom I c= dom Initialized I by GRFUNC_1:8; assume m < n; then insloc m in dom I by A1,SCMFSA_4:def 4; hence insloc m in dom Initialized I by A2; end; hence thesis by SCMFSA_4:def 4; end; end; Lm1: now let s be State of SCM+FSA; let I be Macro-Instruction; A1: ProgramPart Initialized I = I by Th25; hereby assume A2: Initialized I is_pseudo-closed_on s; set k = pseudo-LifeSpan(s,Initialized I); IC (Computation (s +* (Initialized I +* Start-At insloc 0))).k = insloc card ProgramPart Initialized I & for n being Nat st n < k holds IC (Computation (s +* (Initialized I +* Start-At insloc 0))).n in dom Initialized I by A2,SCMFSA8A:def 5; then IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k = insloc card ProgramPart Initialized I by Th16; then A3: IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k = insloc card ProgramPart I by A1,AMI_5:72; A4: now let n be Nat; assume n < k; then IC (Computation (s +* (Initialized I +* Start-At insloc 0))).n in dom Initialized I by A2,SCMFSA8A:def 5; then IC (Computation (Initialize s +* (I +* Start-At insloc 0))).n in dom Initialized I by Th16; hence IC (Computation (Initialize s +* (I +* Start-At insloc 0))).n in dom I by Th20; end; then A5: for n be Nat st not IC (Computation (Initialize s +* (I +* Start-At insloc 0))) .n in dom I holds k <= n; thus I is_pseudo-closed_on Initialize s by A3,A4,SCMFSA8A:def 3; hence pseudo-LifeSpan(s,Initialized I) = pseudo-LifeSpan(Initialize s,I) by A3,A5,SCMFSA8A:def 5; end; assume A6: I is_pseudo-closed_on Initialize s; set k = pseudo-LifeSpan(Initialize s,I); IC (Computation (Initialize s +* (I +* Start-At insloc 0))).k = insloc card ProgramPart I & for n being Nat st n < k holds IC (Computation (Initialize s +* (I +* Start-At insloc 0))).n in dom I by A6,SCMFSA8A:def 5; then IC (Computation (s +* (Initialized I +* Start-At insloc 0))).k = insloc card ProgramPart I by Th16; then A7: IC (Computation (s +* (Initialized I +* Start-At insloc 0))).k = insloc card ProgramPart Initialized I by A1,AMI_5:72; A8: now let n be Nat; assume n < k; then IC (Computation (Initialize s +* (I +* Start-At insloc 0))).n in dom I by A6,SCMFSA8A:def 5; then IC (Computation (s +* (Initialized I +* Start-At insloc 0))).n in dom I by Th16; hence IC (Computation (s +* (Initialized I +* Start-At insloc 0))).n in dom Initialized I by Th20; end; then A9: for n be Nat st not IC (Computation (s +* (Initialized I +* Start-At insloc 0))).n in dom Initialized I holds k <= n; thus Initialized I is_pseudo-closed_on s by A7,A8,SCMFSA8A:def 3; hence pseudo-LifeSpan(s,Initialized I) = pseudo-LifeSpan(Initialize s,I) by A7,A9,SCMFSA8A:def 5; end; theorem ::TMP8 for s being State of SCM+FSA, I being Macro-Instruction holds Initialized I is_pseudo-closed_on s iff I is_pseudo-closed_on Initialize s by Lm1; theorem for s being State of SCM+FSA, I being Macro-Instruction st I is_pseudo-closed_on Initialize s holds pseudo-LifeSpan(s,Initialized I) = pseudo-LifeSpan(Initialize s,I) by Lm1; theorem for s being State of SCM+FSA, I being Macro-Instruction st Initialized I is_pseudo-closed_on s holds pseudo-LifeSpan(s,Initialized I) = pseudo-LifeSpan(Initialize s,I) by Lm1; theorem Th50: ::TMP14 for s being State of SCM+FSA, I being initial FinPartState of SCM+FSA st I is_pseudo-closed_on s holds I is_pseudo-closed_on s +* (I +* Start-At insloc 0) & pseudo-LifeSpan(s,I) = pseudo-LifeSpan(s +* (I +* Start-At insloc 0),I) proof let s be State of SCM+FSA; let I be initial FinPartState of SCM+FSA; assume A1: I is_pseudo-closed_on s; set s2 = s +* (I +* Start-At insloc 0) +* (I +* Start-At insloc 0); A2: s +* (I +* Start-At insloc 0) +* (I +* Start-At insloc 0) = s +* ((I +* Start-At insloc 0) +* (I +* Start-At insloc 0)) by FUNCT_4:15 .= s +* (I +* Start-At insloc 0); then A3: IC (Computation s2).pseudo-LifeSpan(s,I) = insloc card ProgramPart I & for n being Nat st not IC (Computation s2).n in dom I holds pseudo-LifeSpan(s,I) <= n by A1,SCMFSA8A:def 5; IC (Computation s2).pseudo-LifeSpan(s,I) = insloc card ProgramPart I & for n being Nat st n < pseudo-LifeSpan(s,I) holds IC (Computation s2).n in dom I by A1,A2,SCMFSA8A:def 5; hence I is_pseudo-closed_on s +* (I +* Start-At insloc 0) by SCMFSA8A:def 3 ; hence pseudo-LifeSpan(s,I) = pseudo-LifeSpan(s +* (I +* Start-At insloc 0),I) by A3,SCMFSA8A:def 5; end; theorem Th51: ::P'115' for s1,s2 being State of SCM+FSA, I being Macro-Instruction st I +* Start-At insloc 0 c= s1 & I is_pseudo-closed_on s1 for n being Nat st ProgramPart Relocated(I,n) c= s2 & IC s2 = insloc n & s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) holds ((for i being Nat st i < pseudo-LifeSpan(s1,I) holds IncAddr(CurInstr (Computation s1).i,n) = CurInstr (Computation s2).i) & for i being Nat st i <= pseudo-LifeSpan(s1,I) holds IC (Computation s1).i + n = IC (Computation s2).i & (Computation s1).i | (Int-Locations \/ FinSeq-Locations) = (Computation s2).i | (Int-Locations \/ FinSeq-Locations)) proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; assume A1: I +* Start-At insloc 0 c= s1; assume A2: I is_pseudo-closed_on s1; let n be Nat; assume A3: ProgramPart Relocated(I,n) c= s2; assume A4: IC s2 = insloc n; assume A5: s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations); set C1 = Computation s1; set C2 = Computation s2; thus A6: now let i be Nat; assume A7: i < pseudo-LifeSpan(s1,I); defpred P[Nat] means $1 < pseudo-LifeSpan(s1,I) implies IC C1.$1 + n = IC C2.$1 & IncAddr(CurInstr (C1.$1),n) = CurInstr (C2.$1) & C1.$1 | (Int-Locations \/ FinSeq-Locations) = C2.$1 | (Int-Locations \/ FinSeq-Locations); A8: ProgramPart Relocated(I,n) = Relocated(I,n) | the Instruction-Locations of SCM+FSA by AMI_5:def 6; A9: P[0] proof assume A10: 0 < pseudo-LifeSpan(s1,I); A11: IC SCM+FSA in dom (I +* Start-At insloc 0) by SF_MASTR:65; IC C1.0 = IC s1 by AMI_1:def 19 .= s1.IC SCM+FSA by AMI_1:def 15 .= (I +* Start-At insloc 0).IC SCM+FSA by A1,A11,GRFUNC_1:8 .= insloc 0 by SF_MASTR:66; hence IC C1.0 + n = insloc (0 + n) by SCMFSA_4:def 1 .= IC C2.0 by A4,AMI_1:def 19; A12: I c= I +* Start-At insloc 0 by SCMFSA8A:9; then A13: dom I c= dom (I +* Start-At insloc 0) by GRFUNC_1:8; IC (Computation (s1 +* (I +* Start-At insloc 0))).0 = IC (s1 +* (I +* Start-At insloc 0)) by AMI_1:def 19 .= IC (s1 +* I +* Start-At insloc 0) by FUNCT_4:15 .= insloc 0 by AMI_5:79; then A14: insloc 0 in dom I by A2,A10,SCMFSA8A:def 5; A15: IC SCM+FSA in dom (I +* Start-At insloc 0) by SF_MASTR:65; A16: s1.IC s1 = s1.(s1.IC SCM+FSA) by AMI_1:def 15 .= s1.((I +* Start-At insloc 0).IC SCM+FSA) by A1,A15,GRFUNC_1:8 .= s1.insloc 0 by SF_MASTR:66 .= (I +* Start-At insloc 0).insloc 0 by A1,A13,A14,GRFUNC_1:8 .= I.insloc 0 by A12,A14,GRFUNC_1:8; insloc 0 + n in dom Relocated(I,n) by A14,SCMFSA_5:4; then insloc 0 + n in dom ProgramPart Relocated(I,n) by AMI_5:73; then A17: insloc (0 + n) in dom ProgramPart Relocated(I,n) by SCMFSA_4:def 1; A18: insloc 0 in dom ProgramPart I by A14,AMI_5:72; thus IncAddr(CurInstr (C1.0),n) = IncAddr(CurInstr s1,n) by AMI_1:def 19 .= IncAddr(s1.IC s1,n) by AMI_1:def 17 .= Relocated(I,n).(insloc 0 + n) by A16,A18,SCMFSA_5:7 .= Relocated(I,n).insloc (0 + n) by SCMFSA_4:def 1 .= (ProgramPart Relocated(I,n)).insloc n by A8,FUNCT_1:72 .= s2.IC s2 by A3,A4,A17,GRFUNC_1:8 .= CurInstr s2 by AMI_1:def 17 .= CurInstr (C2.0) by AMI_1:def 19; thus C1.0 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) by A5,AMI_1:def 19 .= C2.0 | (Int-Locations \/ FinSeq-Locations) by AMI_1:def 19; end; A19: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume A20: P[k]; assume A21: k + 1 < pseudo-LifeSpan(s1,I); A22: k + 0 < k + 1 by REAL_1:53; A23: C1.(k + 1) = Following C1.k by AMI_1:def 19 .= Exec(CurInstr C1.k,C1.k) by AMI_1:def 18; A24: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18; hence A25: IC C1.(k + 1) + n = IC C2.(k + 1) by A20,A21,A22,A23,AXIOMS:22,SCMFSA6A:41; reconsider j = CurInstr C1.(k + 1) as Instruction of SCM+FSA; reconsider l = IC C1.(k + 1) as Instruction-Location of SCM+FSA; A26: I c= I +* Start-At insloc 0 by SCMFSA8A:9; then A27: dom I c= dom (I +* Start-At insloc 0) by GRFUNC_1:8; s1 +* (I +* Start-At insloc 0) = s1 by A1,AMI_5:10; then A28: IC C1.(k + 1) in dom I by A2,A21,SCMFSA8A:def 5; ProgramPart I = I | the Instruction-Locations of SCM+FSA by AMI_5:def 6; then dom ProgramPart I = dom I /\ the Instruction-Locations of SCM+FSA by FUNCT_1:68; then A29: l in dom ProgramPart I by A28,XBOOLE_0:def 3; A30: j = C1.(k + 1).IC C1.(k + 1) by AMI_1:def 17 .= s1.IC C1.(k + 1) by AMI_1:54 .= (I +* Start-At insloc 0).IC C1.(k + 1) by A1,A27,A28,GRFUNC_1:8 .= I.l by A26,A28,GRFUNC_1:8; IC C2.(k + 1) in dom Relocated(I,n) by A25,A28,SCMFSA_5:4; then IC C2.(k + 1) in dom Relocated(I,n) /\ the Instruction-Locations of SCM+FSA by XBOOLE_0:def 3; then A31: IC C2.(k + 1) in dom ProgramPart Relocated(I,n) by A8,FUNCT_1:68; thus IncAddr(CurInstr C1.(k + 1),n) = Relocated(I,n).(l + n) by A29,A30,SCMFSA_5:7 .= (ProgramPart Relocated(I,n)).(IC C2.(k + 1)) by A8,A25,FUNCT_1:72 .= s2.IC C2.(k + 1) by A3,A31,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) | (Int-Locations \/ FinSeq-Locations) = C2.(k + 1) | (Int-Locations \/ FinSeq-Locations) by A20,A21,A22,A23, A24,AXIOMS:22,SCMFSA6A:41; end; for k being Nat holds P[k] from Ind(A9,A19); hence IncAddr(CurInstr ((Computation s1).i),n) = CurInstr ((Computation s2).i) by A7; end; let i be Nat; assume A32: i <= pseudo-LifeSpan(s1,I); defpred P[Nat] means $1 <= pseudo-LifeSpan(s1,I) implies IC C1.$1 + n = IC C2.$1 & C1.$1 | (Int-Locations \/ FinSeq-Locations) = C2.$1 | (Int-Locations \/ FinSeq-Locations); A33: P[0] proof assume 0 <= pseudo-LifeSpan(s1,I); A34: IC SCM+FSA in dom (I +* Start-At insloc 0) by SF_MASTR:65; IC C1.0 = IC s1 by AMI_1:def 19 .= s1.IC SCM+FSA by AMI_1:def 15 .= (I +* Start-At insloc 0).IC SCM+FSA by A1,A34,GRFUNC_1:8 .= insloc 0 by SF_MASTR:66; hence IC C1.0 + n = insloc (0 + n) by SCMFSA_4:def 1 .= IC C2.0 by A4,AMI_1:def 19; thus C1.0 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) by A5,AMI_1:def 19 .= C2.0 | (Int-Locations \/ FinSeq-Locations) by AMI_1:def 19; end; A35: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume A36: P[k]; assume A37: k + 1 <= pseudo-LifeSpan(s1,I); then A38: k + 1 <= pseudo-LifeSpan(s1,I) + 1 by NAT_1:37; A39: k < pseudo-LifeSpan(s1,I) by A37,NAT_1:38; set i = CurInstr C1.k; A40: C1.(k + 1) = Following C1.k by AMI_1:def 19 .= Exec(CurInstr C1.k,C1.k) by AMI_1:def 18; A41: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18; thus IC C1.(k + 1) + n = IC Exec(IncAddr(i,n),C2.k) by A36,A38,A40,REAL_1:53,SCMFSA6A:41 .= IC C2.(k + 1) by A6,A39,A41; thus C1.(k + 1) | (Int-Locations \/ FinSeq-Locations) = Exec(IncAddr(i,n),C2.k) | (Int-Locations \/ FinSeq-Locations) by A36,A38,A40,REAL_1:53,SCMFSA6A:41 .= C2.(k + 1) | (Int-Locations \/ FinSeq-Locations) by A6,A39,A41; end; for k being Nat holds P[k] from Ind(A33,A35); hence thesis by A32; end; theorem Th52: ::TMP11 for s1,s2 being State of SCM+FSA, I being Macro-Instruction st s1 | (Int-Locations \/ FinSeq-Locations) = s2 | (Int-Locations \/ FinSeq-Locations) holds I is_pseudo-closed_on s1 implies I is_pseudo-closed_on s2 proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; set C1 = Computation (s1 +* (I +* Start-At insloc 0)); set C2 = Computation (s2 +* (I +* Start-At insloc 0)); assume s1 | D = s2 | D; then A1: (s1 +* (I +* Start-At insloc 0)) | D = s2 | D by SCMFSA8A:11 .= (s2 +* (I +* Start-At insloc 0)) | D by SCMFSA8A:11; assume A2: I is_pseudo-closed_on s1; then A3: IC C1.pseudo-LifeSpan(s1,I) = insloc card ProgramPart I & for n being Nat st n < pseudo-LifeSpan(s1,I) holds IC C1.n in dom I by SCMFSA8A:def 5; A4: I is_pseudo-closed_on s1 +* (I +* Start-At insloc 0) by A2,Th50; A5: I +* Start-At insloc 0 c= s1 +* (I +* Start-At insloc 0) by FUNCT_4:26; A6: I +* Start-At insloc 0 c= s2 +* (I +* Start-At insloc 0) by FUNCT_4:26; I c= I +* Start-At insloc 0 by SCMFSA8A:9; then I c= s2 +* (I +* Start-At insloc 0) by A6,XBOOLE_1:1; then A7: ProgramPart Relocated(I,0) c= s2 +* (I +* Start-At insloc 0) by Th9; A8: IC (s2 +* (I +* Start-At insloc 0)) = IC (s2 +* I +* Start-At insloc 0) by FUNCT_4:15 .= insloc 0 by AMI_5:79; A9: IC C2.pseudo-LifeSpan(s1,I) = IC C2.pseudo-LifeSpan(s1 +* (I +* Start-At insloc 0),I) by A2,Th50 .= IC C1.pseudo-LifeSpan(s1 +* (I +* Start-At insloc 0),I) + 0 by A1,A4,A5,A7,A8,Th51 .= IC C1.pseudo-LifeSpan(s1,I) + 0 by A2,Th50 .= IC C1.pseudo-LifeSpan(s1,I) by Th7; now let k be Nat; assume A10: k < pseudo-LifeSpan(s1,I); then k <= pseudo-LifeSpan(s1 +* (I +* Start-At insloc 0),I) by A2,Th50; then IC C2.k = IC C1.k + 0 by A1,A4,A5,A7,A8,Th51 .= IC C1.k by Th7; hence IC C2.k in dom I by A2,A10,SCMFSA8A:def 5; end; hence I is_pseudo-closed_on s2 by A3,A9,SCMFSA8A:def 3; end; theorem Th53: ::T1702 for s being State of SCM+FSA, I being Macro-Instruction st s.intloc 0 = 1 holds I is_pseudo-closed_on s iff I is_pseudo-closed_on Initialize s proof let s be State of SCM+FSA; let I be Macro-Instruction; assume s.intloc 0 = 1; then s | D = (Initialize s) | D by Th27; hence I is_pseudo-closed_on s iff I is_pseudo-closed_on Initialize s by Th52; end; Lm2: for l being Instruction-Location of SCM+FSA holds goto l <> halt SCM+FSA proof let l be Instruction-Location of SCM+FSA; thus thesis by SCMFSA_2:47,124; end; Lm3: for a being Int-Location, l being Instruction-Location of SCM+FSA holds a =0_goto l <> halt SCM+FSA proof let a be Int-Location; let l be Instruction-Location of SCM+FSA; thus thesis by SCMFSA_2:48,124; end; Lm4: for a being Int-Location, l being Instruction-Location of SCM+FSA holds a >0_goto l <> halt SCM+FSA proof let a be Int-Location; let l be Instruction-Location of SCM+FSA; thus thesis by SCMFSA_2:49,124; end; Lm5: for I,J being Macro-Instruction holds ProgramPart Relocated(J,card I) c= I ';' J proof let I,J be Macro-Instruction; I ';' J = Directed I +* ProgramPart Relocated(J,card I) by SCMFSA6A:def 4; hence thesis by FUNCT_4:26; end; theorem Th54: ::TD2' ** n.t for a being Int-Location, I,J being Macro-Instruction holds insloc 0 in dom if=0(a,I,J) & insloc 1 in dom if=0(a,I,J) & insloc 0 in dom if>0(a,I,J) & insloc 1 in dom if>0(a,I,J) proof let a be Int-Location; let I,J be Macro-Instruction; set i = a =0_goto insloc (card J + 3); if=0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1 .= i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62 .= i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62 .= i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:66 .= Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:def 5; then A1: dom Macro i c= dom if=0(a,I,J) by SCMFSA6A:56; dom Macro i = {insloc 0, insloc 1} by SCMFSA7B:4; then insloc 0 in dom Macro i & insloc 1 in dom Macro i by TARSKI:def 2; hence insloc 0 in dom if=0(a,I,J) & insloc 1 in dom if=0(a,I,J) by A1; set i = a >0_goto insloc (card J + 3); if>0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2 .= i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62 .= i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62 .= i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:66 .= Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:def 5; then A2: dom Macro i c= dom if>0(a,I,J) by SCMFSA6A:56; dom Macro i = {insloc 0, insloc 1} by SCMFSA7B:4; then insloc 0 in dom Macro i & insloc 1 in dom Macro i by TARSKI:def 2; hence insloc 0 in dom if>0(a,I,J) & insloc 1 in dom if>0(a,I,J) by A2; end; theorem Th55: ::TD2 ** n.t for a being Int-Location, I,J being Macro-Instruction holds if=0(a,I,J).insloc 0 = a =0_goto insloc (card J + 3) & if=0(a,I,J).insloc 1 = goto insloc 2 & if>0(a,I,J).insloc 0 = a >0_goto insloc (card J + 3) & if>0(a,I,J).insloc 1 = goto insloc 2 proof let a be Int-Location; let I,J be Macro-Instruction; set i = a =0_goto insloc (card J + 3); A1: i <> halt SCM+FSA by Lm3; A2: if=0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1 .= i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62 .= i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62 .= i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:66 .= Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:def 5; dom Macro i = {insloc 0, insloc 1} by SCMFSA7B:4; then A3: insloc 0 in dom Macro i & insloc 1 in dom Macro i by TARSKI:def 2; hence if=0(a,I,J).insloc 0 = (Directed Macro i).insloc 0 by A2,SCMFSA8A:28 .= i by A1,SCMFSA7B:7; thus if=0(a,I,J).insloc 1 = (Directed Macro i).insloc 1 by A2,A3,SCMFSA8A:28 .= goto insloc 2 by SCMFSA7B:8; set i = a >0_goto insloc (card J + 3); A4: i <> halt SCM+FSA by Lm4; A5: if>0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2 .= i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62 .= i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62 .= i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:66 .= Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:def 5; dom Macro i = {insloc 0, insloc 1} by SCMFSA7B:4; then A6: insloc 0 in dom Macro i & insloc 1 in dom Macro i by TARSKI:def 2; hence if>0(a,I,J).insloc 0 = (Directed Macro i).insloc 0 by A5,SCMFSA8A:28 .= i by A4,SCMFSA7B:7; thus if>0(a,I,J).insloc 1 = (Directed Macro i).insloc 1 by A5,A6,SCMFSA8A:28 .= goto insloc 2 by SCMFSA7B:8; end; theorem Th56: ::T6327 ** n.t for a being Int-Location, I,J being Macro-Instruction, n being Nat st n < card I + card J + 3 holds insloc n in dom if=0(a,I,J) & if=0(a,I,J).insloc n <> halt SCM+FSA proof let a be Int-Location; let I,J be Macro-Instruction; let n be Nat; assume A1: n < card I + card J + 3; A2: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; set J1 = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; card J1 = card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a =0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then insloc n in dom J1 by A1,SCMFSA6A:15; then A3: insloc n in dom Directed J1 by SCMFSA6A:14; Directed J1 c= if=0(a,I,J) by A2,SCMFSA6A:55; then A4: dom Directed J1 c= dom if=0(a,I,J) & if=0(a,I,J).insloc n = (Directed J1).insloc n by A3,GRFUNC_1:8; hence insloc n in dom if=0(a,I,J) by A3; (Directed J1).insloc n in rng Directed J1 by A3,FUNCT_1:def 5; hence if=0(a,I,J).insloc n <> halt SCM+FSA by A4,SCMFSA7B:def 6; end; theorem Th57: ::T6327' ** n.t for a being Int-Location, I,J being Macro-Instruction, n being Nat st n < card I + card J + 3 holds insloc n in dom if>0(a,I,J) & if>0(a,I,J).insloc n <> halt SCM+FSA proof let a be Int-Location; let I,J be Macro-Instruction; let n be Nat; assume A1: n < card I + card J + 3; A2: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; set J1 = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; card J1 = card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a >0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then insloc n in dom J1 by A1,SCMFSA6A:15; then A3: insloc n in dom Directed J1 by SCMFSA6A:14; Directed J1 c= if>0(a,I,J) by A2,SCMFSA6A:55; then A4: dom Directed J1 c= dom if>0(a,I,J) & if>0(a,I,J).insloc n = (Directed J1).insloc n by A3,GRFUNC_1:8; hence insloc n in dom if>0(a,I,J) by A3; (Directed J1).insloc n in rng Directed J1 by A3,FUNCT_1:def 5; hence if>0(a,I,J).insloc n <> halt SCM+FSA by A4,SCMFSA7B:def 6; end; theorem Th58: ::T29502 for s being State of SCM+FSA, I being Macro-Instruction st Directed I is_pseudo-closed_on s holds I ';' SCM+FSA-Stop is_closed_on s & I ';' SCM+FSA-Stop is_halting_on s & LifeSpan (s +* (I ';' SCM+FSA-Stop +* Start-At insloc 0)) = pseudo-LifeSpan(s,Directed I) & (for n being Nat st n < pseudo-LifeSpan(s,Directed I) holds IC (Computation (s +* (I +* Start-At insloc 0))).n = IC (Computation (s +* (I ';' SCM+FSA-Stop +* Start-At insloc 0))).n) & for n being Nat st n <= pseudo-LifeSpan(s,Directed I) holds (Computation (s +* (I +* Start-At insloc 0))).n | D = (Computation (s +* (I ';' SCM+FSA-Stop +* Start-At insloc 0))).n | D proof let s be State of SCM+FSA; let I be Macro-Instruction; set I0 = Directed I; set I1 = I ';' SCM+FSA-Stop; set s00 = s +* (I0 +* Start-At insloc 0); set s10 = s +* (I1 +* Start-At insloc 0); assume A1: I0 is_pseudo-closed_on s; set k = pseudo-LifeSpan(s00,I0); A2: I0 c= I1 by SCMFSA6A:55; then A3: dom I0 c= dom I1 by GRFUNC_1:8; card I1 = card I + 1 by SCMFSA6A:61,SCMFSA8A:17; then card I < card I1 by NAT_1:38; then A4: insloc card I in dom I1 by SCMFSA6A:15; A5: card I < card I + card SCM+FSA-Stop by NAT_1:38,SCMFSA8A:17; halt SCM+FSA = SCM+FSA-Stop.insloc (card I -' card I) by GOBOARD9:1,SCMFSA8A:16; then I1.insloc card I = IncAddr(halt SCM+FSA,card I) by A5,Th13 .= halt SCM+FSA by SCMFSA_4:8; then A6: s10.insloc card I = halt SCM+FSA by A4,Th26; A7: I0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; A8: s00 +* (I0 +* Start-At insloc 0) = s +* ((I0 +* Start-At insloc 0) +* (I0 +* Start-At insloc 0)) by FUNCT_4:15 .= s00; A9: I0 is_pseudo-closed_on s00 by A1,Th50; A10: k = pseudo-LifeSpan(s,I0) by A1,Th50; I1 c= I1 +* Start-At insloc 0 & I1 +* Start-At insloc 0 c= s10 by FUNCT_4:26,SCMFSA8A:9; then ProgramPart Relocated(I0,0) c= I1 & I1 c= s10 by A2,Th9,XBOOLE_1:1; then A11: ProgramPart Relocated(I0,0) c= s10 by XBOOLE_1:1; s10 = s +* I1 +* Start-At insloc 0 by FUNCT_4:15; then A12: IC s10 = insloc 0 by AMI_5:79; A13: s00 | D = s | D by SCMFSA8A:11 .= s10 | D by SCMFSA8A:11; A14: now let n be Nat; assume A15: n < pseudo-LifeSpan(s00,I0); then IncAddr(CurInstr (Computation s00).n,0) = CurInstr (Computation s10 ).n by A7,A9,A11,A12,A13,Th51; hence CurInstr (Computation s00).n = CurInstr (Computation s10).n by Th8; thus IC (Computation s00).n in dom I0 by A8,A9,A15,SCMFSA8A:31; thus CurInstr (Computation s00).n <> halt SCM+FSA by A8,A9,A15,SCMFSA8A: 31; end; A16: now let n be Nat; assume A17: n <= pseudo-LifeSpan(s00,I0); then IC (Computation s00).n + 0 = IC (Computation s10).n by A7,A9,A11,A12,A13,Th51; hence IC (Computation s00).n = IC (Computation s10).n by Th7; thus (Computation s00).n | D = (Computation s10).n | D by A7,A9,A11,A12,A13,A17,Th51; end; defpred P[Nat] means k <= $1 implies IC (Computation s10).$1 = insloc card I & CurInstr (Computation s10).$1 = halt SCM+FSA; A18: P[0] proof assume k <= 0; then k = 0 by NAT_1:19; hence A19: IC (Computation s10).0 = IC (Computation s00).k by A16 .= insloc card ProgramPart I0 by A1,A10,SCMFSA8A:def 5 .= insloc card I0 by AMI_5:72 .= insloc card I by SCMFSA8A:34; thus CurInstr (Computation s10).0 = (Computation s10).0.IC (Computation s10).0 by AMI_1:def 17 .= halt SCM+FSA by A6,A19,AMI_1:def 19; end; A20: for n being Nat st P[n] holds P[n + 1] proof let n be Nat; assume A21: P[n]; assume A22: k <= n + 1; thus A23: now per cases by A22,NAT_1:26; suppose k = n + 1; hence IC (Computation s10).(n + 1) = IC (Computation s00).k by A16 .= insloc card ProgramPart I0 by A1,A10,SCMFSA8A:def 5 .= insloc card I0 by AMI_5:72 .= insloc card I by SCMFSA8A:34; suppose A24: k <= n; (Computation s10).(n + 1) = Following (Computation s10).n by AMI_1:def 19 .= Exec(CurInstr (Computation s10).n,(Computation s10).n) by AMI_1:def 18; hence IC (Computation s10).(n + 1) = insloc card I by A21,A24,AMI_1:def 8; end; thus CurInstr (Computation s10).(n + 1) = (Computation s10).(n + 1).IC (Computation s10).(n + 1) by AMI_1:def 17 .= halt SCM+FSA by A6,A23,AMI_1:54; end; A25: for n being Nat holds P[n] from Ind(A18,A20); now let n be Nat; per cases; suppose A26: n < k; then IC (Computation s00).n = IC (Computation s10).n by A16; then IC (Computation s10).n in dom I0 by A1,A10,A26,SCMFSA8A:def 5; hence IC (Computation s10).n in dom I1 by A3; suppose k <= n; hence IC (Computation s10).n in dom I1 by A4,A25; end; hence I1 is_closed_on s by SCMFSA7B:def 7; P[k] by A25; then A27: s10 is halting by AMI_1:def 20; hence I1 is_halting_on s by SCMFSA7B:def 8; A28: CurInstr (Computation s10).k = halt SCM+FSA by A25; now let n be Nat; assume A29: CurInstr (Computation s10).n = halt SCM+FSA; assume A30: k > n; reconsider l = IC (Computation s00).n as Instruction-Location of SCM+FSA; A31: l in dom I0 by A1,A10,A30,SCMFSA8A:def 5; CurInstr (Computation s10).n = CurInstr (Computation s00).n by A14,A30 .= (Computation s00).n.l by AMI_1:def 17 .= s00.l by AMI_1:54 .= I0.l by A31,Th26; then halt SCM+FSA in rng I0 by A29,A31,FUNCT_1:def 5; hence contradiction by SCMFSA7B:def 6; end; then A32: LifeSpan s10 = k by A27,A28,SCM_1:def 2; IC (Computation s10).k = insloc card I by A25; then A33: IC (Computation s00).LifeSpan s10 = insloc card I by A16,A32; A34: card ProgramPart I0 = card I0 by AMI_5:72 .= card I by SCMFSA8A:34; for n be Nat st not IC (Computation s00).n in dom I0 holds LifeSpan s10 <= n by A14,A32; hence LifeSpan s10 = pseudo-LifeSpan(s,I0) by A1,A33,A34,SCMFSA8A:def 5; set s1 = s +* (I +* Start-At insloc 0); defpred P[Nat] means $1 < pseudo-LifeSpan(s,I0) implies IC (Computation s1).$1 in dom I & IC (Computation s1).$1 = IC (Computation s10).$1 & (Computation s1).$1 | D = (Computation s10).$1 | D; A35: P[0] proof assume 0 < pseudo-LifeSpan(s,I0); then IC (Computation (s +* (I0 +* Start-At insloc 0))).0 in dom I0 by A1,SCMFSA8A:31; then IC (s +* (I0 +* Start-At insloc 0)) in dom I0 by AMI_1:def 19; then A36: insloc 0 in dom I0 by Th31; A37: IC (Computation s1).0 = IC s1 by AMI_1:def 19 .= IC (s +* I +* Start-At insloc 0) by FUNCT_4:15 .= insloc 0 by AMI_5:79; hence IC (Computation s1).0 in dom I by A36,SCMFSA6A:14; thus IC (Computation s1).0 = IC (Computation s10).0 by A12,A37,AMI_1:def 19; thus (Computation s1).0 | D = s1 | D by AMI_1:def 19 .= s | D by SCMFSA8A:11 .= s10 | D by SCMFSA8A:11 .= (Computation s10).0 | D by AMI_1:def 19; end; A38: for n being Nat st P[n] holds P[n + 1] proof let n be Nat; set l = IC (Computation s1).n; set l0 = IC (Computation s10).n; assume A39: P[n]; assume A40: n + 1 < pseudo-LifeSpan(s,I0); A41: pseudo-LifeSpan(s,I0) = k by A1,Th50; n < k by A10,A40,NAT_1:37; then A42: IC (Computation s00).(n + 1) = IC (Computation s10).(n + 1) & (Computation s00).n | D = (Computation s10).n | D & l in dom I & l = l0 & (Computation s1).n | D = (Computation s10).n | D by A16,A39,A40,A41; A43: l0 in dom I0 by A39,A40,NAT_1:37,SCMFSA6A:14; A44: now assume A45: I.l = halt SCM+FSA; A46: l0 in dom I & dom I = dom I0 by A39,A40,NAT_1:37,SCMFSA6A:14; n < k by A10,A40,NAT_1:37; then IC (Computation s00).n = IC (Computation s10).n by A16; then A47: CurInstr (Computation s00).n = (Computation s00).n.l0 by AMI_1: def 17 .= s00.l0 by AMI_1:54 .= I0.l by A39,A40,A46,Th26,NAT_1:37 .= goto insloc card I by A39,A40,A45,NAT_1:37,SCMFSA8A:30; A48: IC (Computation s00).(n + 1) = IC Following (Computation s00).n by AMI_1:def 19 .= IC Exec(goto insloc card I,(Computation s00).n) by A47,AMI_1:def 18 .= Exec(goto insloc card I,(Computation s00).n).IC SCM+FSA by AMI_1:def 15 .= insloc card I by SCMFSA_2:95 .= insloc card I0 by SCMFSA8A:34; IC (Computation s00).(n + 1) in dom I0 by A1,A40,SCMFSA8A:31; hence contradiction by A48,SCMFSA6A:15; end; A49: CurInstr (Computation s1).n = (Computation s1).n.l by AMI_1:def 17 .= s1.l by AMI_1:54 .= I.l by A39,A40,Th26,NAT_1:37 .= I0.l0 by A39,A40,A44,NAT_1:37,SCMFSA8A:30 .= I1.l0 by A2,A43,GRFUNC_1:8 .= s10.l0 by A3,A43,Th26 .= (Computation s10).n.l0 by AMI_1:54 .= CurInstr (Computation s10).n by AMI_1:def 17; A50: (Computation s1).(n + 1) = Following (Computation s1).n by AMI_1:def 19 .= Exec(CurInstr (Computation s1).n,(Computation s1).n) by AMI_1:def 18; A51: (Computation s10).(n + 1) = Following (Computation s10).n by AMI_1:def 19 .= Exec(CurInstr (Computation s1).n,(Computation s10).n) by A49,AMI_1:def 18; (for a being Int-Location holds (Computation s1).n.a = (Computation s10).n.a) & for f being FinSeq-Location holds (Computation s1).n.f = (Computation s10).n.f by A39,A40,NAT_1:37,SCMFSA6A:38; then (Computation s1).n,(Computation s10).n equal_outside A by A39,A40,NAT_1:37,SCMFSA6A:28; then A52: (Computation s1).(n + 1),(Computation s10).(n + 1) equal_outside A by A50,A51,SCMFSA6A:32; dom I0 = dom I & IC (Computation s00).(n + 1) in dom I0 by A1,A40,SCMFSA6A:14,SCMFSA8A:31; hence IC (Computation s1).(n + 1) in dom I by A42,A52,SCMFSA6A:29; thus IC (Computation s1).(n + 1) = IC (Computation s10).(n + 1) by A52,SCMFSA6A:29; (for a being Int-Location holds (Computation s1).(n + 1).a = (Computation s10).(n + 1).a) & for f being FinSeq-Location holds (Computation s1).(n + 1).f = (Computation s10).(n + 1).f by A52,SCMFSA6A:30,31; hence (Computation s1).(n + 1) | D = (Computation s10).(n + 1) | D by SCMFSA6A:38; end; A53: for n being Nat holds P[n] from Ind(A35,A38); hence for n be Nat st n < pseudo-LifeSpan(s,I0) holds IC (Computation s1).n = IC (Computation s10).n; let n be Nat; assume A54: n <= pseudo-LifeSpan(s,Directed I); per cases by A54,REAL_1:def 5; suppose n < pseudo-LifeSpan(s,I0); hence (Computation s1).n | D = (Computation s10).n | D by A53; suppose A55: n = pseudo-LifeSpan(s,I0); hereby per cases by NAT_1:22; suppose A56: n = 0; hence (Computation s1).n | D = s1 | D by AMI_1:def 19 .= s | D by SCMFSA8A:11 .= s10 | D by SCMFSA8A:11 .= (Computation s10).n | D by A56,AMI_1:def 19; suppose ex m st n = m + 1; then consider m being Nat such that A57: n = m + 1; A58: (Computation s1).n = Following (Computation s1).m by A57,AMI_1:def 19 .= Exec(CurInstr (Computation s1).m,(Computation s1).m) by AMI_1:def 18; A59: (Computation s10).n = Following (Computation s10).m by A57,AMI_1:def 19 .= Exec(CurInstr (Computation s10).m,(Computation s10).m) by AMI_1:def 18; set i = CurInstr (Computation s1).m; set l = IC (Computation s1).m; set l0 = IC (Computation s10).m; A60: m + 0 < pseudo-LifeSpan(s,I0) by A55,A57,REAL_1:53; then A61: (Computation s1).m | D = (Computation s10).m | D & l in dom I & l = l0 by A53; then A62: l0 in dom I0 by SCMFSA6A:14; A63: i = (Computation s1).m.l by AMI_1:def 17 .= s1.l by AMI_1:54 .= I.l by A61,Th26; I0 c= I1 by SCMFSA6A:55; then A64: I0.l0 = I1.l0 by A62,GRFUNC_1:8 .= s10.l0 by A3,A62,Th26 .= (Computation s10).m.l0 by AMI_1:54 .= CurInstr (Computation s10).m by AMI_1:def 17; hereby per cases; suppose A65: i = halt SCM+FSA; then CurInstr (Computation s10).m = goto insloc card I by A61,A63,A64,SCMFSA8A:30; then InsCode CurInstr (Computation s10).m = 6 by SCMFSA_2:47; then A66: InsCode CurInstr (Computation s10).m in {0,6,7,8} by ENUMSET1:19; thus (Computation s1).n | D = (Computation s1).m | D by A58,A65,AMI_1:def 8 .= (Computation s10).m | D by A53,A60 .= (Computation s10).n | D by A59,A66,Th32; suppose i <> halt SCM+FSA; then CurInstr (Computation s10).m = i by A61,A63,A64,SCMFSA8A:30; hence (Computation s1).n | D = (Computation s10).n | D by A58,A59,A61,SCMFSA6C:5; end; end; end; theorem Th59: ::T29502' for s being State of SCM+FSA, I being Macro-Instruction st Directed I is_pseudo-closed_on s holds (Result (s +* (I ';' SCM+FSA-Stop +* Start-At insloc 0))) | D = (Computation (s +* (I +* Start-At insloc 0))). pseudo-LifeSpan(s,Directed I) | D proof let s be State of SCM+FSA; let I be Macro-Instruction; set I0 = Directed I; set I1 = I ';' SCM+FSA-Stop; set s2 = s +* (I +* Start-At insloc 0); set s10 = s +* (I1 +* Start-At insloc 0); set k = pseudo-LifeSpan(s,I0); assume I0 is_pseudo-closed_on s; then A1: I1 is_halting_on s & LifeSpan s10 = k & (Computation s2).k | D = (Computation s10).k | D by Th58; then s10 is halting by SCMFSA7B:def 8; hence (Result s10) | D = (Computation s2).k | D by A1,SCMFSA6B:16; end; theorem for s being State of SCM+FSA, I being Macro-Instruction st s.intloc 0 = 1 & Directed I is_pseudo-closed_on s holds IExec(I ';' SCM+FSA-Stop,s) | D = (Computation (s +* (I +* Start-At insloc 0))). pseudo-LifeSpan(s,Directed I) | D proof let s be State of SCM+FSA; let I be Macro-Instruction; set I0 = Directed I; set I1 = I ';' SCM+FSA-Stop; set s2 = s +* (I +* Start-At insloc 0); set s10 = s +* (I1 +* Start-At insloc 0); set k = pseudo-LifeSpan(s,I0); assume A1: s.intloc 0 = 1; assume A2: I0 is_pseudo-closed_on s; thus IExec(I1,s) | D = (Result (s +* Initialized I1) +* s | A) | D by SCMFSA6B:def 1 .= (Result (s +* Initialized I1)) | D by Th35 .= (Result s10) | D by A1,Th18 .= (Computation s2).k | D by A2,Th59; end; theorem Th61: ::TMP20 ** n.t for I,J being Macro-Instruction,a being Int-Location holds if=0(a,I,J).insloc (card I + card J + 3) = halt SCM+FSA proof let I,J be Macro-Instruction; let a be Int-Location; A1: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; set II = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; A2: card II = card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a =0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then A3: card I + card J + 3 < card II + card SCM+FSA-Stop by NAT_1:38,SCMFSA8A :17; card I + card J + 3 -' card II = 0 by A2,GOBOARD9:1; hence if=0(a,I,J).insloc (card I + card J + 3) = IncAddr(halt SCM+FSA,card II) by A1,A2,A3,Th13,SCMFSA8A:16 .= halt SCM+FSA by SCMFSA_4:8; end; theorem Th62: ::TMP20' ** n.t for I,J being Macro-Instruction,a being Int-Location holds if>0(a,I,J).insloc (card I + card J + 3) = halt SCM+FSA proof let I,J be Macro-Instruction; let a be Int-Location; A1: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; set II = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; A2: card II = card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a >0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then A3: card I + card J + 3 < card II + card SCM+FSA-Stop by NAT_1:38,SCMFSA8A :17; card I + card J + 3 -' card II = 0 by A2,GOBOARD9:1; hence if>0(a,I,J).insloc (card I + card J + 3) = IncAddr(halt SCM+FSA,card II) by A1,A2,A3,Th13,SCMFSA8A:16 .= halt SCM+FSA by SCMFSA_4:8; end; theorem Th63: ::TMP21 ** n.t for I,J being Macro-Instruction,a being Int-Location holds if=0(a,I,J).insloc (card J + 2) = goto insloc (card I + card J + 3) proof let I,J be Macro-Instruction; let a be Int-Location; set JJ = a =0_goto insloc (card J + 3) ';' J; set J3 = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1); A1: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; A2: card JJ = card (Macro (a =0_goto insloc (card J + 3)) ';' J) by SCMFSA6A:def 5 .= card Macro (a =0_goto insloc (card J + 3)) + card J by SCMFSA6A:61 .= 2 + card J by SCMFSA7B:6; card Goto insloc (card I + 1) = 1 by SCMFSA8A:29; then A3: card J3 = card J + 2 + 1 by A2,SCMFSA6A:61 .= card J + (2 + 1) by XCMPLX_1:1; card J + 2 -' card JJ = 0 by A2,GOBOARD9:1; then A4: goto insloc (card I + 1) = (Goto insloc (card I + 1)).insloc (card J + 2 -' card JJ) by SCMFSA8A:47; card Goto insloc (card I + 1) = 1 by SCMFSA8A:29; then card J + 2 < card JJ + card Goto insloc (card I + 1) by A2,NAT_1:38; then A5: J3.insloc (card J + 2) = IncAddr(goto insloc (card I + 1),card JJ) by A2,A4,Th13 .= goto (insloc (card I + 1) + (card J + 2)) by A2,SCMFSA_4:14 .= goto insloc (card I + 1 + (card J + 2)) by SCMFSA_4:def 1 .= goto insloc (card I + 1 + card J + 2) by XCMPLX_1:1 .= goto insloc (card I + (card J + 1) + 2) by XCMPLX_1:1 .= goto insloc (card I + (card J + 1 + 2)) by XCMPLX_1:1 .= goto insloc (card I + (card J + (1 + 2))) by XCMPLX_1:1 .= goto insloc (card I + card J + (1 + 2)) by XCMPLX_1:1; A6: 6 <> 0 & InsCode goto insloc (card I + card J + 3) = 6 & InsCode halt SCM+FSA = 0 by SCMFSA_2:47,124; card J3 = card J + 2 + 1 by A3,XCMPLX_1:1; then card J + 2 < card J3 by NAT_1:38; then A7: insloc (card J + 2) in dom J3 by SCMFSA6A:15; then (J3 ';' (I ';' SCM+FSA-Stop)).insloc (card J + 2) = (Directed J3).insloc (card J + 2) by SCMFSA8A:28 .= goto insloc (card I + card J + 3) by A5,A6,A7,SCMFSA8A:30; hence if=0(a,I,J).insloc (card J + 2) = goto insloc (card I + card J + 3) by A1,SCMFSA6A:62; end; theorem Th64: ::TMP21' ** n.t for I,J being Macro-Instruction,a being Int-Location holds if>0(a,I,J).insloc (card J + 2) = goto insloc (card I + card J + 3) proof let I,J be Macro-Instruction; let a be Int-Location; set JJ = a >0_goto insloc (card J + 3) ';' J; set J3 = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1); A1: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; A2: card JJ = card (Macro (a >0_goto insloc (card J + 3)) ';' J) by SCMFSA6A:def 5 .= card Macro (a >0_goto insloc (card J + 3)) + card J by SCMFSA6A:61 .= 2 + card J by SCMFSA7B:6; card Goto insloc (card I + 1) = 1 by SCMFSA8A:29; then A3: card J3 = card J + 2 + 1 by A2,SCMFSA6A:61 .= card J + (2 + 1) by XCMPLX_1:1; card J + 2 -' card JJ = 0 by A2,GOBOARD9:1; then A4: goto insloc (card I + 1) = (Goto insloc (card I + 1)).insloc (card J + 2 -' card JJ) by SCMFSA8A:47; card Goto insloc (card I + 1) = 1 by SCMFSA8A:29; then card J + 2 < card JJ + card Goto insloc (card I + 1) by A2,NAT_1:38; then A5: J3.insloc (card J + 2) = IncAddr(goto insloc (card I + 1),card JJ) by A2,A4,Th13 .= goto (insloc (card I + 1) + (card J + 2)) by A2,SCMFSA_4:14 .= goto insloc (card I + 1 + (card J + 2)) by SCMFSA_4:def 1 .= goto insloc (card I + 1 + card J + 2) by XCMPLX_1:1 .= goto insloc (card I + (card J + 1) + 2) by XCMPLX_1:1 .= goto insloc (card I + (card J + 1 + 2)) by XCMPLX_1:1 .= goto insloc (card I + (card J + (1 + 2))) by XCMPLX_1:1 .= goto insloc (card I + card J + (1 + 2)) by XCMPLX_1:1; A6: 6 <> 0 & InsCode goto insloc (card I + card J + 3) = 6 & InsCode halt SCM+FSA = 0 by SCMFSA_2:47,124; card J3 = card J + 2 + 1 by A3,XCMPLX_1:1; then card J + 2 < card J3 by NAT_1:38; then A7: insloc (card J + 2) in dom J3 by SCMFSA6A:15; then (J3 ';' (I ';' SCM+FSA-Stop)).insloc (card J + 2) = (Directed J3).insloc (card J + 2) by SCMFSA8A:28 .= goto insloc (card I + card J + 3) by A5,A6,A7,SCMFSA8A:30; hence if>0(a,I,J).insloc (card J + 2) = goto insloc (card I + card J + 3) by A1,SCMFSA6A:62; end; theorem Th65: ::T31139 ** n.t for J being Macro-Instruction,a being Int-Location holds if=0(a,Goto insloc 2,J).insloc (card J + 3) = goto insloc (card J + 5) proof let J be Macro-Instruction; let a be Int-Location; set JJ = a =0_goto insloc (card J + 3) ';' J; set J3 = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc 2; set J4 = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc 2 ';' Goto insloc 2; card Goto insloc 2 = 1 by SCMFSA8A:29; then A1: if=0(a,Goto insloc 2,J) = (a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (1 + 1) ';' Goto insloc 2) ';' SCM+FSA-Stop by SCMFSA8B:def 1; A2: card Goto insloc 2 = 1 by SCMFSA8A:29; card JJ = card (Macro (a =0_goto insloc (card J + 3)) ';' J) by SCMFSA6A:def 5 .= card Macro (a =0_goto insloc (card J + 3)) + card J by SCMFSA6A:61 .= 2 + card J by SCMFSA7B:6; then A3: card J3 = card J + 2 + 1 by A2,SCMFSA6A:61 .= card J + (2 + 1) by XCMPLX_1:1; then A4: card J4 = card J + 3 + 1 by A2,SCMFSA6A:61 .= card J + (3 + 1) by XCMPLX_1:1; card J + 3 -' card J3 = 0 by A3,GOBOARD9:1; then A5: goto insloc 2 = (Goto insloc 2).insloc (card J + 3 -' card J3) by SCMFSA8A:47; card Goto insloc 2 = 1 by SCMFSA8A:29; then card J + 3 < card J3 + card Goto insloc 2 by A3,NAT_1:38; then A6: J4.insloc (card J + 3) = IncAddr(goto insloc 2,card J3) by A3,A5,Th13 .= goto (insloc 2 + (card J + 3)) by A3,SCMFSA_4:14 .= goto insloc (2 + (card J + 3)) by SCMFSA_4:def 1 .= goto insloc (card J + (2 + 3)) by XCMPLX_1:1; A7: 6 <> 0 & InsCode goto insloc (card J + 5) = 6 & InsCode halt SCM+FSA = 0 by SCMFSA_2:47,124; card J4 = card J + 3 + 1 by A4,XCMPLX_1:1; then card J + 3 < card J4 by NAT_1:38; then A8: insloc (card J + 3) in dom J4 by SCMFSA6A:15; then (J4 ';' SCM+FSA-Stop).insloc (card J + 3) = (Directed J4).insloc (card J + 3) by SCMFSA8A:28 .= goto insloc (card J + 5) by A6,A7,A8,SCMFSA8A:30; hence if=0(a,Goto insloc 2,J).insloc (card J + 3) = goto insloc (card J + 5) by A1; end; theorem Th66: ::TMP71 for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.a = 0 & Directed I is_pseudo-closed_on s holds if=0(a,I,J) is_halting_on s & if=0(a,I,J) is_closed_on s & LifeSpan (s +* (if=0(a,I,J) +* Start-At insloc 0)) = LifeSpan (s +* (I ';' SCM+FSA-Stop +* Start-At insloc 0)) + 1 proof let s be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set I0 = Directed I; set I1 = I ';' SCM+FSA-Stop; set s00 = s +* (I0 +* Start-At insloc 0); set s3 = s +* (if=0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set C3 = Computation s3; set i = a =0_goto insloc (card J + 3); assume A1: s.a = 0; assume A2: I0 is_pseudo-closed_on s; A3: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; A4: I0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A5: I0 is_pseudo-closed_on s00 by A2,Th52; A6: insloc 0 in dom if=0(a,I,J) by Th54; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A7: dom if=0(a,I,J) c= dom (if=0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A8: IC SCM+FSA in dom (if=0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A9: now thus IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if=0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A8,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; end; now thus s3.insloc 0 = (if=0(a,I,J) +* Start-At insloc 0).insloc 0 by A6,A7,FUNCT_4:14 .= if=0(a,I,J).insloc 0 by A6,SCMFSA6B:7 .= i by Th55; end; then A10: CurInstr s3 = i by A9,AMI_1:def 17; A11: now thus (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; end; A12: now thus card (i ';' J ';' Goto insloc (card I + 1)) = card (Macro i ';' J ';' Goto insloc (card I + 1)) by SCMFSA6A:def 5 .= card (Macro i ';' J) + card Goto insloc (card I + 1) by SCMFSA6A:61 .= card (Macro i ';' J) + 1 by SCMFSA8A:29 .= card Macro i + card J + 1 by SCMFSA6A:61 .= card J + 2 + 1 by SCMFSA7B:6 .= card J + (2 + 1) by XCMPLX_1:1; end; not a in dom (if=0(a,I,J) +* Start-At insloc 0) & a in dom s by SCMFSA6B:12,SCMFSA_2:66; then A13: s3.a = 0 by A1,FUNCT_4:12; A14: if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A15: if=0(a,I,J) c= s3 by A14,XBOOLE_1:1; if=0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I1 by A3,SCMFSA6A:62; then ProgramPart Relocated(I1,card J + 3) c= if=0(a,I,J) by A12,Lm5; then ProgramPart Relocated(I1,card J + 3) c= s3 by A15,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:64; then A16: ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:72; I0 c= I1 by SCMFSA6A:55; then ProgramPart Relocated(I0,card J + 3) c= ProgramPart Relocated(I1,card J + 3) by Th12; then A17: ProgramPart Relocated(I0,card J + 3) c= s4 by A16,XBOOLE_1:1; A18: now thus IC C3.1 = C3.1.IC SCM+FSA by AMI_1:def 15 .= insloc (card J + 3) by A11,A13,SCMFSA_2:96; end; s00,s3 equal_outside A by SCMFSA8A:14; then A19: s00 | D = s3 | D by SCMFSA6A:39; A20: now let a be Int-Location; thus s00.a = s3.a by A19,SCMFSA6A:38 .= s4.a by A11,SCMFSA_2:96; end; now let f be FinSeq-Location; thus s00.f = s3.f by A19,SCMFSA6A:38 .= s4.f by A11,SCMFSA_2:96; end; then A21: s00 | D = s4 | D by A20,SCMFSA6A:38; if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A22: if=0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if=0(a,I,J) by NAT_1:38; then A23: insloc (card I + card J + 3) in dom if=0(a,I,J) by SCMFSA6A:15; A24: now thus IC (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = IC (Computation s4).pseudo-LifeSpan(s00,I0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,I0) + (card J + 3) by A4,A5,A17,A18,A21,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,I0) + (card J + 3) by A2,Th50 .= insloc card ProgramPart I0 + (card J + 3) by A2,SCMFSA8A:def 5 .= insloc card I0 + (card J + 3) by AMI_5:72 .= insloc card I + (card J + 3) by SCMFSA8A:34 .= insloc (card I + (card J + 3)) by SCMFSA_4:def 1 .= insloc (card I + card J + 3) by XCMPLX_1:1; end; then A25: CurInstr (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = (Computation s3).(pseudo-LifeSpan(s00,I0) + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if=0(a,I,J).insloc (card I + card J + 3) by A22,A23,GRFUNC_1:8 .= halt SCM+FSA by Th61; then A26: s3 is halting by AMI_1:def 20; hence if=0(a,I,J) is_halting_on s by SCMFSA7B:def 8; now let k be Nat; per cases by NAT_1:19; suppose k = 0; then (Computation s3).k = s3 by AMI_1:def 19; then IC (Computation s3).k = insloc 0 by Th31; hence IC (Computation s3).k in dom if=0(a,I,J) by Th54; suppose A27: 0 < k & k < pseudo-LifeSpan(s00,I0) + 1; then 0 + 1 <= k by INT_1:20; then consider k1 being Nat such that A28: 1 + k1 = k by NAT_1:28; A29: k1 < pseudo-LifeSpan(s00,I0) by A27,A28,AXIOMS:24; then A30: k1 < pseudo-LifeSpan(s,I0) by A2,Th50; A31: IC (Computation s3).k = IC (Computation s4).k1 by A28,AMI_1:51 .= IC (Computation s00).k1 + (card J + 3) by A4,A5,A17,A18,A21,A29,Th51; consider n being Nat such that A32: IC (Computation s00).k1 = insloc n by SCMFSA_2:21; insloc n in dom I0 by A2,A30,A32,SCMFSA8A:31; then n < card I0 by SCMFSA6A:15; then n + (card J + 3) < card I0 + (card J + 3) by REAL_1:53; then n + (card J + 3) < card I + (card J + 3) by SCMFSA8A:34; then A33: n + (card J + 3) < card I + card J + 3 by XCMPLX_1:1; card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; then card I + card J + 3 < card if=0(a,I,J) by REAL_1:69; then n + (card J + 3) < card if=0(a,I,J) by A33,AXIOMS:22; then insloc (n + (card J + 3)) in dom if=0(a,I,J) by SCMFSA6A:15; hence IC (Computation s3).k in dom if=0(a,I,J) by A31,A32,SCMFSA_4:def 1; suppose 0 < k & pseudo-LifeSpan(s00,I0) + 1 <= k; hence IC (Computation s3).k in dom if=0(a,I,J) by A23,A24,A25,AMI_1:52; end; hence if=0(a,I,J) is_closed_on s by SCMFSA7B:def 7; now let k be Nat; assume A34: CurInstr (Computation s3).k = halt SCM+FSA; assume not pseudo-LifeSpan(s00,I0) + 1 <= k; then A35: k <= pseudo-LifeSpan(s00,I0) by NAT_1:38; A36: IC s3 = insloc 0 by Th31; A37: insloc 0 in dom if=0(a,I,J) by Th54; A38: InsCode halt SCM+FSA = 0 & InsCode (a =0_goto insloc (card J + 3)) = 7 by SCMFSA_2:48,124; CurInstr (Computation s3).0 = CurInstr s3 by AMI_1:def 19 .= s3.insloc 0 by A36,AMI_1:def 17 .= if=0(a,I,J).insloc 0 by A37,Th26 .= a =0_goto insloc (card J + 3) by Th55; then consider k1 being Nat such that A39: k1 + 1 = k by A34,A38,NAT_1:22; k1 < k by A39,REAL_1:69; then A40: k1 < pseudo-LifeSpan(s00,I0) by A35,AXIOMS:22; A41: IC (Computation s3).k = IC (Computation s4).k1 by A39,AMI_1:51 .= IC (Computation s00).k1 + (card J + 3) by A4,A5,A17,A18,A21,A40,Th51; consider n being Nat such that A42: IC (Computation s00).k1 = insloc n by SCMFSA_2:21; k1 < pseudo-LifeSpan(s,I0) by A2,A40,Th50; then insloc n in dom I0 by A2,A42,SCMFSA8A:31; then n < card I0 by SCMFSA6A:15; then n + (card J + 3) < card I0 + (card J + 3) by REAL_1:53; then n + (card J + 3) < card I + (card J + 3) by SCMFSA8A:34; then A43: n + (card J + 3) < card I + card J + 3 by XCMPLX_1:1; card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; then card I + card J + 3 < card if=0(a,I,J) by REAL_1:69; then n + (card J + 3) < card if=0(a,I,J) by A43,AXIOMS:22; then insloc (n + (card J + 3)) in dom if=0(a,I,J) by SCMFSA6A:15; then A44: IC (Computation s3).k in dom if=0(a,I,J) by A41,A42,SCMFSA_4:def 1; set J1 = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; card J1 = card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a =0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then insloc (n + (card J + 3)) in dom J1 by A43,SCMFSA6A:15; then IC (Computation s3).k in dom J1 by A41,A42,SCMFSA_4:def 1; then A45: IC (Computation s3).k in dom Directed J1 by SCMFSA6A:14; Directed J1 c= if=0(a,I,J) by A3,SCMFSA6A:55; then A46: if=0(a,I,J).IC (Computation s3).k = (Directed J1).IC (Computation s3).k by A45,GRFUNC_1:8; A47: (Directed J1).IC (Computation s3).k in rng Directed J1 by A45,FUNCT_1:def 5; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if=0(a,I,J).IC (Computation s3).k by A44,Th26; hence contradiction by A34,A46,A47,SCMFSA7B:def 6; end; then A48: LifeSpan s3 = pseudo-LifeSpan(s00,I0) + 1 by A25,A26,SCM_1:def 2; pseudo-LifeSpan(s,I0) = LifeSpan (s +* (I1 +* Start-At insloc 0)) by A2,Th58; hence LifeSpan s3 = LifeSpan (s +* (I1 +* Start-At insloc 0)) + 1 by A2,A48,Th50; end; theorem for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.intloc 0 = 1 & s.a = 0 & Directed I is_pseudo-closed_on s holds IExec(if=0(a,I,J),s) | (Int-Locations \/ FinSeq-Locations) = IExec(I ';' SCM+FSA-Stop,s) | (Int-Locations \/ FinSeq-Locations) proof let ss be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set I0 = Directed I; set s = Initialize ss; set I1 = I ';' SCM+FSA-Stop; set s00 = s +* (I0 +* Start-At insloc 0); set s3 = s +* (if=0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set C3 = Computation s3; set i = a =0_goto insloc (card J + 3); assume A1: ss.intloc 0 = 1; assume ss.a = 0; then A2: s.a = 0 by SCMFSA6C:3; assume I0 is_pseudo-closed_on ss; then A3: I0 is_pseudo-closed_on s by A1,Th53; A4: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; A5: I0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A6: I0 is_pseudo-closed_on s00 by A3,Th52; A7: insloc 0 in dom if=0(a,I,J) by Th54; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A8: dom if=0(a,I,J) c= dom (if=0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A9: IC SCM+FSA in dom (if=0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A10: now thus IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if=0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A9,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; end; now thus s3.insloc 0 = (if=0(a,I,J) +* Start-At insloc 0).insloc 0 by A7,A8,FUNCT_4:14 .= if=0(a,I,J).insloc 0 by A7,SCMFSA6B:7 .= i by Th55; end; then A11: CurInstr s3 = i by A10,AMI_1:def 17; A12: now thus (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; end; A13: now thus card (i ';' J ';' Goto insloc (card I + 1)) = card (Macro i ';' J ';' Goto insloc (card I + 1)) by SCMFSA6A:def 5 .= card (Macro i ';' J) + card Goto insloc (card I + 1) by SCMFSA6A:61 .= card (Macro i ';' J) + 1 by SCMFSA8A:29 .= card Macro i + card J + 1 by SCMFSA6A:61 .= card J + 2 + 1 by SCMFSA7B:6 .= card J + (2 + 1) by XCMPLX_1:1; end; not a in dom (if=0(a,I,J) +* Start-At insloc 0) & a in dom s by SCMFSA6B:12,SCMFSA_2:66; then A14: s3.a = 0 by A2,FUNCT_4:12; A15: if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A16: if=0(a,I,J) c= s3 by A15,XBOOLE_1:1; if=0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I1 by A4,SCMFSA6A:62; then ProgramPart Relocated(I1,card J + 3) c= if=0(a,I,J) by A13,Lm5; then ProgramPart Relocated(I1,card J + 3) c= s3 by A16,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:64; then A17: ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:72; I0 c= I1 by SCMFSA6A:55; then ProgramPart Relocated(I0,card J + 3) c= ProgramPart Relocated(I1,card J + 3) by Th12; then A18: ProgramPart Relocated(I0,card J + 3) c= s4 by A17,XBOOLE_1:1; A19: now thus IC C3.1 = C3.1.IC SCM+FSA by AMI_1:def 15 .= insloc (card J + 3) by A12,A14,SCMFSA_2:96; end; s00,s3 equal_outside A by SCMFSA8A:14; then A20: s00 | D = s3 | D by SCMFSA6A:39; A21: now let a be Int-Location; thus s00.a = s3.a by A20,SCMFSA6A:38 .= s4.a by A12,SCMFSA_2:96; end; now let f be FinSeq-Location; thus s00.f = s3.f by A20,SCMFSA6A:38 .= s4.f by A12,SCMFSA_2:96; end; then A22: s00 | D = s4 | D by A21,SCMFSA6A:38; if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A23: if=0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if=0(a,I,J) by NAT_1:38; then A24: insloc (card I + card J + 3) in dom if=0(a,I,J) by SCMFSA6A:15; now thus IC (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = IC (Computation s4).pseudo-LifeSpan(s00,I0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,I0) + (card J + 3) by A5,A6,A18,A19,A22,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,I0) + (card J + 3) by A3,Th50 .= insloc card ProgramPart I0 + (card J + 3) by A3,SCMFSA8A:def 5 .= insloc card I0 + (card J + 3) by AMI_5:72 .= insloc card I + (card J + 3) by SCMFSA8A:34 .= insloc (card I + (card J + 3)) by SCMFSA_4:def 1 .= insloc (card I + card J + 3) by XCMPLX_1:1; end; then A25: CurInstr (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = (Computation s3).(pseudo-LifeSpan(s00,I0) + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if=0(a,I,J).insloc (card I + card J + 3) by A23,A24,GRFUNC_1:8 .= halt SCM+FSA by Th61; then A26: s3 is halting by AMI_1:def 20; now let k be Nat; assume A27: CurInstr (Computation s3).k = halt SCM+FSA; assume not pseudo-LifeSpan(s00,I0) + 1 <= k; then A28: k <= pseudo-LifeSpan(s00,I0) by NAT_1:38; A29: IC s3 = insloc 0 by Th31; A30: insloc 0 in dom if=0(a,I,J) by Th54; A31: InsCode halt SCM+FSA = 0 & InsCode (a =0_goto insloc (card J + 3)) = 7 by SCMFSA_2:48,124; CurInstr (Computation s3).0 = CurInstr s3 by AMI_1:def 19 .= s3.insloc 0 by A29,AMI_1:def 17 .= if=0(a,I,J).insloc 0 by A30,Th26 .= a =0_goto insloc (card J + 3) by Th55; then consider k1 being Nat such that A32: k1 + 1 = k by A27,A31,NAT_1:22; k1 < k by A32,REAL_1:69; then A33: k1 < pseudo-LifeSpan(s00,I0) by A28,AXIOMS:22; A34: IC (Computation s3).k = IC (Computation s4).k1 by A32,AMI_1:51 .= IC (Computation s00).k1 + (card J + 3) by A5,A6,A18,A19,A22,A33,Th51; consider n being Nat such that A35: IC (Computation s00).k1 = insloc n by SCMFSA_2:21; k1 < pseudo-LifeSpan(s,I0) by A3,A33,Th50; then insloc n in dom I0 by A3,A35,SCMFSA8A:31; then n < card I0 by SCMFSA6A:15; then n + (card J + 3) < card I0 + (card J + 3) by REAL_1:53; then n + (card J + 3) < card I + (card J + 3) by SCMFSA8A:34; then A36: n + (card J + 3) < card I + card J + 3 by XCMPLX_1:1; card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; then card I + card J + 3 < card if=0(a,I,J) by REAL_1:69; then n + (card J + 3) < card if=0(a,I,J) by A36,AXIOMS:22; then insloc (n + (card J + 3)) in dom if=0(a,I,J) by SCMFSA6A:15; then A37: IC (Computation s3).k in dom if=0(a,I,J) by A34,A35,SCMFSA_4:def 1; set J1 = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; card J1 = card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a =0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a =0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then insloc (n + (card J + 3)) in dom J1 by A36,SCMFSA6A:15; then IC (Computation s3).k in dom J1 by A34,A35,SCMFSA_4:def 1; then A38: IC (Computation s3).k in dom Directed J1 by SCMFSA6A:14; Directed J1 c= if=0(a,I,J) by A4,SCMFSA6A:55; then A39: if=0(a,I,J).IC (Computation s3).k = (Directed J1).IC (Computation s3).k by A38,GRFUNC_1:8; A40: (Directed J1).IC (Computation s3).k in rng Directed J1 by A38,FUNCT_1:def 5; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if=0(a,I,J).IC (Computation s3).k by A37,Th26; hence contradiction by A27,A39,A40,SCMFSA7B:def 6; end; then A41: LifeSpan s3 = pseudo-LifeSpan(s00,I0) + 1 by A25,A26,SCM_1:def 2; set s1 = s +* (I1 +* Start-At insloc 0); s +* Initialized if=0(a,I,J) = Initialize s +* (if=0(a,I,J) +* Start-At insloc 0) & s +* Initialized I1 = Initialize s +* (I1 +* Start-At insloc 0) by SCMFSA8A:13; then A42: s +* Initialized if=0(a,I,J) = s3 & s +* Initialized I1 = s1 by Th15; A43: I1 is_halting_on s & LifeSpan s1 = pseudo-LifeSpan(s,I0) by A3,Th58; then A44: s1 is halting by SCMFSA7B:def 8; A45: Directed I0 = I0 by SCMFSA8A:40; I0 ';' SCM+FSA-Stop = I1 & Directed I0 is_pseudo-closed_on s by A3,SCMFSA8A:40,41; then A46: (Computation s00).pseudo-LifeSpan(s,I0) | D = (Computation s1).pseudo-LifeSpan(s,I0) | D by A45,Th58; thus IExec(if=0(a,I,J),ss) | D = IExec(if=0(a,I,J),s) | D by Th17 .= (Result (s +* Initialized if=0(a,I,J)) +* s | A) | D by SCMFSA6B:def 1 .= (Result s3) | D by A42,Th35 .= (Computation s3).(LifeSpan s3) | D by A26,SCMFSA6B:16 .= (Computation s4).pseudo-LifeSpan(s00,I0) | D by A41,AMI_1:51 .= (Computation s00).pseudo-LifeSpan(s00,I0) | D by A5,A6,A18,A19,A22,Th51 .= (Computation s1).(LifeSpan s1) | D by A3,A43,A46,Th50 .= (Result s1) | D by A44,SCMFSA6B:16 .= (Result (s +* Initialized I1) +* s | A) | D by A42,Th35 .= IExec(I1,s) | D by SCMFSA6B:def 1 .= IExec(I1,ss) | D by Th17; end; theorem Th68: ::TMP71' for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.a > 0 & Directed I is_pseudo-closed_on s holds if>0(a,I,J) is_halting_on s & if>0(a,I,J) is_closed_on s & LifeSpan (s +* (if>0(a,I,J) +* Start-At insloc 0)) = LifeSpan (s +* (I ';' SCM+FSA-Stop +* Start-At insloc 0)) + 1 proof let s be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set I0 = Directed I; set I1 = I ';' SCM+FSA-Stop; set s00 = s +* (I0 +* Start-At insloc 0); set s3 = s +* (if>0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set C3 = Computation s3; set i = a >0_goto insloc (card J + 3); assume A1: s.a > 0; assume A2: I0 is_pseudo-closed_on s; A3: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; A4: I0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A5: I0 is_pseudo-closed_on s00 by A2,Th52; A6: insloc 0 in dom if>0(a,I,J) by Th54; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A7: dom if>0(a,I,J) c= dom (if>0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A8: IC SCM+FSA in dom (if>0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A9: now thus IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if>0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A8,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; end; now thus s3.insloc 0 = (if>0(a,I,J) +* Start-At insloc 0).insloc 0 by A6,A7,FUNCT_4:14 .= if>0(a,I,J).insloc 0 by A6,SCMFSA6B:7 .= i by Th55; end; then A10: CurInstr s3 = i by A9,AMI_1:def 17; A11: now thus (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; end; A12: now thus card (i ';' J ';' Goto insloc (card I + 1)) = card (Macro i ';' J ';' Goto insloc (card I + 1)) by SCMFSA6A:def 5 .= card (Macro i ';' J) + card Goto insloc (card I + 1) by SCMFSA6A:61 .= card (Macro i ';' J) + 1 by SCMFSA8A:29 .= card Macro i + card J + 1 by SCMFSA6A:61 .= card J + 2 + 1 by SCMFSA7B:6 .= card J + (2 + 1) by XCMPLX_1:1; end; not a in dom (if>0(a,I,J) +* Start-At insloc 0) & a in dom s by SCMFSA6B:12,SCMFSA_2:66; then A13: s3.a = s.a by FUNCT_4:12; A14: if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A15: if>0(a,I,J) c= s3 by A14,XBOOLE_1:1; if>0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I1 by A3,SCMFSA6A:62; then ProgramPart Relocated(I1,card J + 3) c= if>0(a,I,J) by A12,Lm5; then ProgramPart Relocated(I1,card J + 3) c= s3 by A15,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:64; then A16: ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:72; I0 c= I1 by SCMFSA6A:55; then ProgramPart Relocated(I0,card J + 3) c= ProgramPart Relocated(I1,card J + 3) by Th12; then A17: ProgramPart Relocated(I0,card J + 3) c= s4 by A16,XBOOLE_1:1; A18: now thus IC C3.1 = C3.1.IC SCM+FSA by AMI_1:def 15 .= insloc (card J + 3) by A1,A11,A13,SCMFSA_2:97; end; s00,s3 equal_outside A by SCMFSA8A:14; then A19: s00 | D = s3 | D by SCMFSA6A:39; A20: now let a be Int-Location; thus s00.a = s3.a by A19,SCMFSA6A:38 .= s4.a by A11,SCMFSA_2:97; end; now let f be FinSeq-Location; thus s00.f = s3.f by A19,SCMFSA6A:38 .= s4.f by A11,SCMFSA_2:97; end; then A21: s00 | D = s4 | D by A20,SCMFSA6A:38; if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A22: if>0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if>0(a,I,J) by NAT_1:38; then A23: insloc (card I + card J + 3) in dom if>0(a,I,J) by SCMFSA6A:15; A24: now thus IC (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = IC (Computation s4).pseudo-LifeSpan(s00,I0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,I0) + (card J + 3) by A4,A5,A17,A18,A21,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,I0) + (card J + 3) by A2,Th50 .= insloc card ProgramPart I0 + (card J + 3) by A2,SCMFSA8A:def 5 .= insloc card I0 + (card J + 3) by AMI_5:72 .= insloc card I + (card J + 3) by SCMFSA8A:34 .= insloc (card I + (card J + 3)) by SCMFSA_4:def 1 .= insloc (card I + card J + 3) by XCMPLX_1:1; end; then A25: CurInstr (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = (Computation s3).(pseudo-LifeSpan(s00,I0) + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if>0(a,I,J).insloc (card I + card J + 3) by A22,A23,GRFUNC_1:8 .= halt SCM+FSA by Th62; then A26: s3 is halting by AMI_1:def 20; hence if>0(a,I,J) is_halting_on s by SCMFSA7B:def 8; now let k be Nat; per cases by NAT_1:19; suppose k = 0; then (Computation s3).k = s3 by AMI_1:def 19; then IC (Computation s3).k = insloc 0 by Th31; hence IC (Computation s3).k in dom if>0(a,I,J) by Th54; suppose A27: 0 < k & k < pseudo-LifeSpan(s00,I0) + 1; then 0 + 1 <= k by INT_1:20; then consider k1 being Nat such that A28: 1 + k1 = k by NAT_1:28; A29: k1 < pseudo-LifeSpan(s00,I0) by A27,A28,AXIOMS:24; then A30: k1 < pseudo-LifeSpan(s,I0) by A2,Th50; A31: IC (Computation s3).k = IC (Computation s4).k1 by A28,AMI_1:51 .= IC (Computation s00).k1 + (card J + 3) by A4,A5,A17,A18,A21,A29,Th51; consider n being Nat such that A32: IC (Computation s00).k1 = insloc n by SCMFSA_2:21; insloc n in dom I0 by A2,A30,A32,SCMFSA8A:31; then n < card I0 by SCMFSA6A:15; then n + (card J + 3) < card I0 + (card J + 3) by REAL_1:53; then n + (card J + 3) < card I + (card J + 3) by SCMFSA8A:34; then A33: n + (card J + 3) < card I + card J + 3 by XCMPLX_1:1; card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; then card I + card J + 3 < card if>0(a,I,J) by REAL_1:69; then n + (card J + 3) < card if>0(a,I,J) by A33,AXIOMS:22; then insloc (n + (card J + 3)) in dom if>0(a,I,J) by SCMFSA6A:15; hence IC (Computation s3).k in dom if>0(a,I,J) by A31,A32,SCMFSA_4:def 1 ; suppose 0 < k & pseudo-LifeSpan(s00,I0) + 1 <= k; hence IC (Computation s3).k in dom if>0(a,I,J) by A23,A24,A25,AMI_1:52; end; hence if>0(a,I,J) is_closed_on s by SCMFSA7B:def 7; now let k be Nat; assume A34: CurInstr (Computation s3).k = halt SCM+FSA; assume not pseudo-LifeSpan(s00,I0) + 1 <= k; then A35: k <= pseudo-LifeSpan(s00,I0) by NAT_1:38; A36: IC s3 = insloc 0 by Th31; A37: insloc 0 in dom if>0(a,I,J) by Th54; A38: InsCode halt SCM+FSA = 0 & InsCode (a >0_goto insloc (card J + 3)) = 8 by SCMFSA_2:49,124; CurInstr (Computation s3).0 = CurInstr s3 by AMI_1:def 19 .= s3.insloc 0 by A36,AMI_1:def 17 .= if>0(a,I,J).insloc 0 by A37,Th26 .= a >0_goto insloc (card J + 3) by Th55; then consider k1 being Nat such that A39: k1 + 1 = k by A34,A38,NAT_1:22; k1 < k by A39,REAL_1:69; then A40: k1 < pseudo-LifeSpan(s00,I0) by A35,AXIOMS:22; A41: IC (Computation s3).k = IC (Computation s4).k1 by A39,AMI_1:51 .= IC (Computation s00).k1 + (card J + 3) by A4,A5,A17,A18,A21,A40,Th51; consider n being Nat such that A42: IC (Computation s00).k1 = insloc n by SCMFSA_2:21; k1 < pseudo-LifeSpan(s,I0) by A2,A40,Th50; then insloc n in dom I0 by A2,A42,SCMFSA8A:31; then n < card I0 by SCMFSA6A:15; then n + (card J + 3) < card I0 + (card J + 3) by REAL_1:53; then n + (card J + 3) < card I + (card J + 3) by SCMFSA8A:34; then A43: n + (card J + 3) < card I + card J + 3 by XCMPLX_1:1; card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; then card I + card J + 3 < card if>0(a,I,J) by REAL_1:69; then n + (card J + 3) < card if>0(a,I,J) by A43,AXIOMS:22; then insloc (n + (card J + 3)) in dom if>0(a,I,J) by SCMFSA6A:15; then A44: IC (Computation s3).k in dom if>0(a,I,J) by A41,A42,SCMFSA_4:def 1; set J1 = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; card J1 = card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a >0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then insloc (n + (card J + 3)) in dom J1 by A43,SCMFSA6A:15; then IC (Computation s3).k in dom J1 by A41,A42,SCMFSA_4:def 1; then A45: IC (Computation s3).k in dom Directed J1 by SCMFSA6A:14; Directed J1 c= if>0(a,I,J) by A3,SCMFSA6A:55; then A46: if>0(a,I,J).IC (Computation s3).k = (Directed J1).IC (Computation s3).k by A45,GRFUNC_1:8; A47: (Directed J1).IC (Computation s3).k in rng Directed J1 by A45,FUNCT_1:def 5; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if>0(a,I,J).IC (Computation s3).k by A44,Th26; hence contradiction by A34,A46,A47,SCMFSA7B:def 6; end; then A48: LifeSpan s3 = pseudo-LifeSpan(s00,I0) + 1 by A25,A26,SCM_1:def 2; pseudo-LifeSpan(s,I0) = LifeSpan (s +* (I1 +* Start-At insloc 0)) by A2,Th58; hence LifeSpan s3 = LifeSpan (s +* (I1 +* Start-At insloc 0)) + 1 by A2,A48,Th50; end; theorem Th69: ::TMP7' for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.intloc 0 = 1 & s.a > 0 & Directed I is_pseudo-closed_on s holds IExec(if>0(a,I,J),s) | (Int-Locations \/ FinSeq-Locations) = IExec(I ';' SCM+FSA-Stop,s) | (Int-Locations \/ FinSeq-Locations) proof let ss be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set I0 = Directed I; set s = Initialize ss; set I1 = I ';' SCM+FSA-Stop; set s00 = s +* (I0 +* Start-At insloc 0); set s3 = s +* (if>0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set C3 = Computation s3; set i = a >0_goto insloc (card J + 3); assume A1: ss.intloc 0 = 1; assume ss.a > 0; then A2: s.a > 0 by SCMFSA6C:3; assume I0 is_pseudo-closed_on ss; then A3: I0 is_pseudo-closed_on s by A1,Th53; A4: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; A5: I0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A6: I0 is_pseudo-closed_on s00 by A3,Th52; A7: insloc 0 in dom if>0(a,I,J) by Th54; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A8: dom if>0(a,I,J) c= dom (if>0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A9: IC SCM+FSA in dom (if>0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A10: now thus IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if>0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A9,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; end; now thus s3.insloc 0 = (if>0(a,I,J) +* Start-At insloc 0).insloc 0 by A7,A8,FUNCT_4:14 .= if>0(a,I,J).insloc 0 by A7,SCMFSA6B:7 .= i by Th55; end; then A11: CurInstr s3 = i by A10,AMI_1:def 17; A12: now thus (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; end; A13: now thus card (i ';' J ';' Goto insloc (card I + 1)) = card (Macro i ';' J ';' Goto insloc (card I + 1)) by SCMFSA6A:def 5 .= card (Macro i ';' J) + card Goto insloc (card I + 1) by SCMFSA6A:61 .= card (Macro i ';' J) + 1 by SCMFSA8A:29 .= card Macro i + card J + 1 by SCMFSA6A:61 .= card J + 2 + 1 by SCMFSA7B:6 .= card J + (2 + 1) by XCMPLX_1:1; end; not a in dom (if>0(a,I,J) +* Start-At insloc 0) & a in dom s by SCMFSA6B:12,SCMFSA_2:66; then A14: s3.a = s.a by FUNCT_4:12; A15: if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A16: if>0(a,I,J) c= s3 by A15,XBOOLE_1:1; if>0(a,I,J) = i ';' J ';' Goto insloc (card I + 1) ';' I1 by A4,SCMFSA6A:62; then ProgramPart Relocated(I1,card J + 3) c= if>0(a,I,J) by A13,Lm5; then ProgramPart Relocated(I1,card J + 3) c= s3 by A16,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:64; then A17: ProgramPart Relocated(I1,card J + 3) c= s4 by AMI_5:72; I0 c= I1 by SCMFSA6A:55; then ProgramPart Relocated(I0,card J + 3) c= ProgramPart Relocated(I1,card J + 3) by Th12; then A18: ProgramPart Relocated(I0,card J + 3) c= s4 by A17,XBOOLE_1:1; A19: now thus IC C3.1 = C3.1.IC SCM+FSA by AMI_1:def 15 .= insloc (card J + 3) by A2,A12,A14,SCMFSA_2:97; end; s00,s3 equal_outside A by SCMFSA8A:14; then A20: s00 | D = s3 | D by SCMFSA6A:39; A21: now let a be Int-Location; thus s00.a = s3.a by A20,SCMFSA6A:38 .= s4.a by A12,SCMFSA_2:97; end; now let f be FinSeq-Location; thus s00.f = s3.f by A20,SCMFSA6A:38 .= s4.f by A12,SCMFSA_2:97; end; then A22: s00 | D = s4 | D by A21,SCMFSA6A:38; if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A23: if>0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if>0(a,I,J) by NAT_1:38; then A24: insloc (card I + card J + 3) in dom if>0(a,I,J) by SCMFSA6A:15; now thus IC (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = IC (Computation s4).pseudo-LifeSpan(s00,I0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,I0) + (card J + 3) by A5,A6,A18,A19,A22,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,I0) + (card J + 3) by A3,Th50 .= insloc card ProgramPart I0 + (card J + 3) by A3,SCMFSA8A:def 5 .= insloc card I0 + (card J + 3) by AMI_5:72 .= insloc card I + (card J + 3) by SCMFSA8A:34 .= insloc (card I + (card J + 3)) by SCMFSA_4:def 1 .= insloc (card I + card J + 3) by XCMPLX_1:1; end; then A25: CurInstr (Computation s3).(pseudo-LifeSpan(s00,I0) + 1) = (Computation s3).(pseudo-LifeSpan(s00,I0) + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if>0(a,I,J).insloc (card I + card J + 3) by A23,A24,GRFUNC_1:8 .= halt SCM+FSA by Th62; then A26: s3 is halting by AMI_1:def 20; now let k be Nat; assume A27: CurInstr (Computation s3).k = halt SCM+FSA; assume not pseudo-LifeSpan(s00,I0) + 1 <= k; then A28: k <= pseudo-LifeSpan(s00,I0) by NAT_1:38; A29: IC s3 = insloc 0 by Th31; A30: insloc 0 in dom if>0(a,I,J) by Th54; A31: InsCode halt SCM+FSA = 0 & InsCode (a >0_goto insloc (card J + 3)) = 8 by SCMFSA_2:49,124; CurInstr (Computation s3).0 = CurInstr s3 by AMI_1:def 19 .= s3.insloc 0 by A29,AMI_1:def 17 .= if>0(a,I,J).insloc 0 by A30,Th26 .= a >0_goto insloc (card J + 3) by Th55; then consider k1 being Nat such that A32: k1 + 1 = k by A27,A31,NAT_1:22; k1 < k by A32,REAL_1:69; then A33: k1 < pseudo-LifeSpan(s00,I0) by A28,AXIOMS:22; A34: IC (Computation s3).k = IC (Computation s4).k1 by A32,AMI_1:51 .= IC (Computation s00).k1 + (card J + 3) by A5,A6,A18,A19,A22,A33,Th51; consider n being Nat such that A35: IC (Computation s00).k1 = insloc n by SCMFSA_2:21; k1 < pseudo-LifeSpan(s,I0) by A3,A33,Th50; then insloc n in dom I0 by A3,A35,SCMFSA8A:31; then n < card I0 by SCMFSA6A:15; then n + (card J + 3) < card I0 + (card J + 3) by REAL_1:53; then n + (card J + 3) < card I + (card J + 3) by SCMFSA8A:34; then A36: n + (card J + 3) < card I + card J + 3 by XCMPLX_1:1; card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; then card I + card J + 3 < card if>0(a,I,J) by REAL_1:69; then n + (card J + 3) < card if>0(a,I,J) by A36,AXIOMS:22; then insloc (n + (card J + 3)) in dom if>0(a,I,J) by SCMFSA6A:15; then A37: IC (Computation s3).k in dom if>0(a,I,J) by A34,A35,SCMFSA_4:def 1; set J1 = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I; card J1 = card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1) ';' I) by SCMFSA6A:def 5 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J ';' Goto insloc (card I + 1)) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + card Goto insloc (card I + 1) + card I by SCMFSA6A:61 .= card (Macro (a >0_goto insloc (card J + 3)) ';' J) + 1 + card I by SCMFSA8A:29 .= card Macro (a >0_goto insloc (card J + 3)) + card J + 1 + card I by SCMFSA6A:61 .= 2 + card J + 1 + card I by SCMFSA7B:6 .= card J + (2 + 1) + card I by XCMPLX_1:1 .= card I + card J + 3 by XCMPLX_1:1; then insloc (n + (card J + 3)) in dom J1 by A36,SCMFSA6A:15; then IC (Computation s3).k in dom J1 by A34,A35,SCMFSA_4:def 1; then A38: IC (Computation s3).k in dom Directed J1 by SCMFSA6A:14; Directed J1 c= if>0(a,I,J) by A4,SCMFSA6A:55; then A39: if>0(a,I,J).IC (Computation s3).k = (Directed J1).IC (Computation s3).k by A38,GRFUNC_1:8; A40: (Directed J1).IC (Computation s3).k in rng Directed J1 by A38,FUNCT_1:def 5; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if>0(a,I,J).IC (Computation s3).k by A37,Th26; hence contradiction by A27,A39,A40,SCMFSA7B:def 6; end; then A41: LifeSpan s3 = pseudo-LifeSpan(s00,I0) + 1 by A25,A26,SCM_1:def 2; set s1 = s +* (I1 +* Start-At insloc 0); s +* Initialized if>0(a,I,J) = Initialize s +* (if>0(a,I,J) +* Start-At insloc 0) & s +* Initialized I1 = Initialize s +* (I1 +* Start-At insloc 0) by SCMFSA8A:13; then A42: s +* Initialized if>0(a,I,J) = s3 & s +* Initialized I1 = s1 by Th15; A43: I1 is_halting_on s & LifeSpan s1 = pseudo-LifeSpan(s,I0) by A3,Th58; then A44: s1 is halting by SCMFSA7B:def 8; A45: Directed I0 = I0 by SCMFSA8A:40; I0 ';' SCM+FSA-Stop = I1 & Directed I0 is_pseudo-closed_on s by A3,SCMFSA8A:40,41; then A46: (Computation s00).pseudo-LifeSpan(s,I0) | D = (Computation s1).pseudo-LifeSpan(s,I0) | D by A45,Th58; thus IExec(if>0(a,I,J),ss) | D = IExec(if>0(a,I,J),s) | D by Th17 .= (Result (s +* Initialized if>0(a,I,J)) +* s | A) | D by SCMFSA6B:def 1 .= (Result s3) | D by A42,Th35 .= (Computation s3).(LifeSpan s3) | D by A26,SCMFSA6B:16 .= (Computation s4).pseudo-LifeSpan(s00,I0) | D by A41,AMI_1:51 .= (Computation s00).pseudo-LifeSpan(s00,I0) | D by A5,A6,A18,A19,A22,Th51 .= (Computation s1).(LifeSpan s1) | D by A3,A43,A46,Th50 .= (Result s1) | D by A44,SCMFSA6B:16 .= (Result (s +* Initialized I1) +* s | A) | D by A42,Th35 .= IExec(I1,s) | D by SCMFSA6B:def 1 .= IExec(I1,ss) | D by Th17; end; theorem Th70: ::TMP171 for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.a <> 0 & Directed J is_pseudo-closed_on s holds if=0(a,I,J) is_halting_on s & if=0(a,I,J) is_closed_on s & LifeSpan (s +* (if=0(a,I,J) +* Start-At insloc 0)) = LifeSpan (s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0)) + 3 proof let s be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set J0 = Directed J; set s0 = Initialize s; set J9 = J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)); set s00 = s +* (J0 +* Start-At insloc 0); set s3 = s +* (if=0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set s5 = (Computation s3).2; set i = a =0_goto insloc (card J + 3); assume s.a <> 0; then A1: s0.a <> 0 by SCMFSA6C:3; A2: s3.a = s.a by Th28 .= s0.a by SCMFSA6C:3; assume A3: J0 is_pseudo-closed_on s; A4: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; A5: J0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A6: J0 is_pseudo-closed_on s00 by A3,Th52; A7: insloc 0 in dom if=0(a,I,J) by Th54; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A8: dom if=0(a,I,J) c= dom (if=0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A9: IC SCM+FSA in dom (if=0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A10: IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if=0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A9,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; s3.insloc 0 = (if=0(a,I,J) +* Start-At insloc 0).insloc 0 by A7,A8,FUNCT_4 :14 .= if=0(a,I,J).insloc 0 by A7,SCMFSA6B:7 .= i by Th55; then A11: CurInstr s3 = i by A10,AMI_1:def 17; A12: now thus (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; end; A13: now thus IC s4 = s4.IC SCM+FSA by AMI_1:def 15 .= Next IC s3 by A1,A2,A12,SCMFSA_2:96 .= insloc (0 + 1) by A10,SCMFSA_2:32; end; A14: insloc 1 in dom if=0(a,I,J) by Th54; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A15: dom if=0(a,I,J) c= dom (if=0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; now thus s4.insloc 1 = s3.insloc 1 by AMI_1:54 .= (if=0(a,I,J) +* Start-At insloc 0).insloc 1 by A14,A15,FUNCT_4:14 .= if=0(a,I,J).insloc 1 by A14,SCMFSA6B:7 .= goto insloc 2 by Th55; end; then A16: CurInstr s4 = goto insloc 2 by A13,AMI_1:def 17; A17: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(goto insloc 2,s4) by A16,AMI_1:def 18; A18: card Macro i = 2 by SCMFSA7B:6; A19: if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A20: if=0(a,I,J) c= s3 by A19,XBOOLE_1:1; if=0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by A4,SCMFSA6A:def 5; then if=0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62; then if=0(a,I,J) = Macro i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62; then if=0(a,I,J) = Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:62; then ProgramPart Relocated(J9,2) c= if=0(a,I,J) by A18,Lm5; then ProgramPart Relocated(J9,2) c= s3 by A20,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(J9,2) c= s5 by AMI_5:64; then A21: ProgramPart Relocated(J9,2) c= s5 by AMI_5:72; J0 c= J9 by SCMFSA6A:55; then ProgramPart Relocated(J0,2) c= ProgramPart Relocated(J9,2) by Th12; then A22: ProgramPart Relocated(J0,2) c= s5 by A21,XBOOLE_1:1; A23: IC s5 = s5.IC SCM+FSA by AMI_1:def 15 .= insloc 2 by A17,SCMFSA_2:95; s00,s3 equal_outside A by SCMFSA8A:14; then A24: s00 | D = s3 | D by SCMFSA6A:39; A25: now let a be Int-Location; thus s00.a = s3.a by A24,SCMFSA6A:38 .= s4.a by A12,SCMFSA_2:96 .= s5.a by A17,SCMFSA_2:95; end; now let f be FinSeq-Location; thus s00.f = s3.f by A24,SCMFSA6A:38 .= s4.f by A12,SCMFSA_2:96 .= s5.f by A17,SCMFSA_2:95; end; then A26: s00 | D = s5 | D by A25,SCMFSA6A:38; if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A27: if=0(a,I,J) c= s3 by SCMFSA6B:5; A28: 0 < card I + 2 by NAT_1:19; now thus card if=0(a,I,J) = card I + card J + (2 + 2) by SCMFSA8B:14 .= card I + card J + 2 + 2 by XCMPLX_1:1 .= card I + (card J + 2) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 2) by XCMPLX_1:1; end; then A29: card J + 2 + 0 < card if=0(a,I,J) by A28,REAL_1:67; then A30: insloc (card J + 2) in dom if=0(a,I,J) by SCMFSA6A:15; set ss = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2); A31: now thus IC ss = IC (Computation s5).pseudo-LifeSpan(s00,J0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,J0) + 2 by A5,A6,A22,A23,A26,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,J0) + 2 by A3,Th50 .= insloc card ProgramPart J0 + 2 by A3,SCMFSA8A:def 5 .= insloc card J0 + 2 by AMI_5:72 .= insloc (card J0 + 2) by SCMFSA_4:def 1 .= insloc (card J + 2) by SCMFSA8A:34; end; then A32: CurInstr ss = ss.insloc (card J + 2) by AMI_1:def 17 .= s3.insloc (card J + 2) by AMI_1:54 .= if=0(a,I,J).insloc (card J + 2) by A27,A30,GRFUNC_1:8 .= goto insloc (card I + card J + 3) by Th63; if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A33: if=0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if=0(a,I,J) by NAT_1:38; then A34: insloc (card I + card J + 3) in dom if=0(a,I,J) by SCMFSA6A:15; A35: now thus IC (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = IC Following ss by AMI_1:def 19 .= IC Exec(goto insloc (card I + card J + 3),ss) by A32,AMI_1:def 18 .= Exec(goto insloc (card I + card J + 3),ss).IC SCM+FSA by AMI_1:def 15 .= insloc (card I + card J + 3) by SCMFSA_2:95; end; then A36: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if=0(a,I,J).insloc (card I + card J + 3) by A33,A34,GRFUNC_1:8 .= halt SCM+FSA by Th61; then A37: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + (2 + 1)) = halt SCM+FSA by XCMPLX_1:1; A38: s3 is halting by A36,AMI_1:def 20; hence if=0(a,I,J) is_halting_on s by SCMFSA7B:def 8; now let k be Nat; per cases by NAT_1:19; suppose k = 0; then (Computation s3).k = s3 by AMI_1:def 19; then IC (Computation s3).k = insloc 0 by Th31; hence IC (Computation s3).k in dom if=0(a,I,J) by Th54; suppose 0 < k & k = 1; hence IC (Computation s3).k in dom if=0(a,I,J) by A13,Th54; suppose A39: 0 < k & k <> 1 & k < pseudo-LifeSpan(s00,J0) + 2; then 0 + 1 <= k by INT_1:20; then 1 < k by A39,REAL_1:def 5; then 0 + (1 + 1) <= k by INT_1:20; then consider k2 being Nat such that A40: 2 + k2 = k by NAT_1:28; A41: k2 < pseudo-LifeSpan(s00,J0) by A39,A40,AXIOMS:24; then A42: k2 < pseudo-LifeSpan(s,J0) by A3,Th50; A43: IC (Computation s3).k = IC (Computation s5).k2 by A40,AMI_1:51 .= IC (Computation s00).k2 + 2 by A5,A6,A22,A23,A26,A41,Th51; consider n being Nat such that A44: IC (Computation s00).k2 = insloc n by SCMFSA_2:21; insloc n in dom J0 by A3,A42,A44,SCMFSA8A:31; then n < card J0 by SCMFSA6A:15; then n + 2 < card J0 + 2 by REAL_1:53; then A45: n + 2 < card J + 2 by SCMFSA8A:34; card if=0(a,I,J) = card I + card J + (2 + 2) by SCMFSA8B:14 .= card I + card J + 2 + 2 by XCMPLX_1:1 .= card J + (card I + 2) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 2) by XCMPLX_1:1; then card J + 2 <= card if=0(a,I,J) by NAT_1:37; then n + 2 < card if=0(a,I,J) by A45,AXIOMS:22; then insloc (n + 2) in dom if=0(a,I,J) by SCMFSA6A:15; hence IC (Computation s3).k in dom if=0(a,I,J) by A43,A44,SCMFSA_4:def 1 ; suppose A46: 0 < k & k <> 1 & pseudo-LifeSpan(s00,J0) + 2 <= k; hereby per cases by A46,REAL_1:def 5; suppose pseudo-LifeSpan(s00,J0) + 2 = k; hence IC (Computation s3).k in dom if=0(a,I,J) by A29,A31,SCMFSA6A:15; suppose pseudo-LifeSpan(s00,J0) + 2 < k; then pseudo-LifeSpan(s00,J0) + 2 + 1 <= k by INT_1:20; hence IC (Computation s3).k in dom if=0(a,I,J) by A34,A35,A36,AMI_1:52 ; end; end; hence if=0(a,I,J) is_closed_on s by SCMFSA7B:def 7; now let k be Nat; assume A47: CurInstr (Computation s3).k = halt SCM+FSA; assume A48: not pseudo-LifeSpan(s00,J0) + (1 + 2) <= k; CurInstr (Computation s3).0 = i by A11,AMI_1:def 19; then A49: k <> 0 & k <> 1 by A16,A47,SCMFSA_2:47,48,124; 2 < 3 & 0 <= card I + card J by NAT_1:18; then 0 + 2 < card I + card J + 3 by REAL_1:67; then A50: insloc 2 in dom if=0(a,I,J) & if=0(a,I,J).insloc 2 <> halt SCM+FSA by Th56; CurInstr (Computation s3).2 = (Computation s3).2.insloc 2 by A23,AMI_1:def 17 .= s3.insloc 2 by AMI_1:54 .= if=0(a,I,J).insloc 2 by A50,Th26; then 2 < k by A47,A49,A50,CQC_THE1:3; then consider k2 being Nat such that A51: 2 + k2 = k by NAT_1:28; k < pseudo-LifeSpan(s00,J0) + 1 + 2 by A48,XCMPLX_1:1; then k2 < pseudo-LifeSpan(s00,J0) + 1 by A51,AXIOMS:24; then A52: k2 <= pseudo-LifeSpan(s00,J0) by NAT_1:38; consider n being Nat such that A53: IC (Computation s00).k2 = insloc n by SCMFSA_2:21; A54: IC (Computation s3).k = IC (Computation s5).k2 by A51,AMI_1:51 .= IC (Computation s00).k2 + 2 by A5,A6,A22,A23,A26,A52,Th51 .= insloc (n + 2) by A53,SCMFSA_4:def 1; A55: k2 <= pseudo-LifeSpan(s,J0) by A3,A52,Th50; A56: now per cases by A55,REAL_1:def 5; suppose k2 = pseudo-LifeSpan(s,J0); then IC (Computation s00).k2 = insloc card ProgramPart J0 by A3,SCMFSA8A:def 5 .= insloc card J0 by AMI_5:72; then A57: n = card J0 by A53,SCMFSA_2:18 .= card J by SCMFSA8A:34; card I + card J + (2 + 1) = card I + (card J + (2 + 1)) by XCMPLX_1: 1 .= card J + 2 + 1 + card I by XCMPLX_1:1; then card J + 2 + 1 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A57,NAT_1:38; suppose k2 < pseudo-LifeSpan(s,J0); then insloc n in dom J0 by A3,A53,SCMFSA8A:31; then n < card J0 by SCMFSA6A:15; then n + 2 < card J0 + 2 by REAL_1:53; then A58: n + 2 < card J + 2 by SCMFSA8A:34; card I + card J + (1 + 2) = card I + card J + 1 + 2 by XCMPLX_1:1 .= card J + (card I + 1) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 1) by XCMPLX_1:1; then card J + 2 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A58,AXIOMS:22; end; then A59: insloc (n + 2) in dom if=0(a,I,J) & if=0(a,I,J).insloc (n + 2) <> halt SCM+FSA by Th56; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if=0(a,I,J).IC (Computation s3).k by A54,A59,Th26; hence contradiction by A47,A54,A56,Th56; end; then A60: LifeSpan s3 = pseudo-LifeSpan(s00,J0) + 3 by A37,A38,SCM_1:def 2; pseudo-LifeSpan(s,J0) = LifeSpan (s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0)) by A3,Th58; hence LifeSpan (s +* (if=0(a,I,J) +* Start-At insloc 0)) = LifeSpan (s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0)) + 3 by A3,A60,Th50; end; theorem for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.intloc 0 = 1 & s.a <> 0 & Directed J is_pseudo-closed_on s holds IExec(if=0(a,I,J),s) | (Int-Locations \/ FinSeq-Locations) = IExec(J ';' SCM+FSA-Stop,s) | (Int-Locations \/ FinSeq-Locations) proof let ss be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set J0 = Directed J; set s = Initialize ss; set s0 = Initialize s; set J9 = J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)); set s00 = s +* (J0 +* Start-At insloc 0); set s3 = s +* (if=0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set s5 = (Computation s3).2; set i = a =0_goto insloc (card J + 3); assume A1: ss.intloc 0 = 1; assume ss.a <> 0; then s.a <> 0 by SCMFSA6C:3; then A2: s0.a <> 0 by SCMFSA6C:3; A3: s3.a = s.a by Th28 .= s0.a by SCMFSA6C:3; assume J0 is_pseudo-closed_on ss; then A4: J0 is_pseudo-closed_on s by A1,Th53; A5: if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; A6: J0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A7: J0 is_pseudo-closed_on s00 by A4,Th52; A8: insloc 0 in dom if=0(a,I,J) by Th54; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A9: dom if=0(a,I,J) c= dom (if=0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A10: IC SCM+FSA in dom (if=0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A11: IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if=0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A10,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; s3.insloc 0 = (if=0(a,I,J) +* Start-At insloc 0).insloc 0 by A8,A9,FUNCT_4 :14 .= if=0(a,I,J).insloc 0 by A8,SCMFSA6B:7 .= i by Th55; then A12: CurInstr s3 = i by A11,AMI_1:def 17; A13: now thus (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A12,AMI_1:def 18; end; A14: now thus IC s4 = s4.IC SCM+FSA by AMI_1:def 15 .= Next IC s3 by A2,A3,A13,SCMFSA_2:96 .= insloc (0 + 1) by A11,SCMFSA_2:32; end; A15: insloc 1 in dom if=0(a,I,J) by Th54; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A16: dom if=0(a,I,J) c= dom (if=0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; now thus s4.insloc 1 = s3.insloc 1 by AMI_1:54 .= (if=0(a,I,J) +* Start-At insloc 0).insloc 1 by A15,A16,FUNCT_4:14 .= if=0(a,I,J).insloc 1 by A15,SCMFSA6B:7 .= goto insloc 2 by Th55; end; then A17: CurInstr s4 = goto insloc 2 by A14,AMI_1:def 17; A18: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(goto insloc 2,s4) by A17,AMI_1:def 18; A19: card Macro i = 2 by SCMFSA7B:6; A20: if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if=0(a,I,J) c= if=0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A21: if=0(a,I,J) c= s3 by A20,XBOOLE_1:1; if=0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by A5,SCMFSA6A:def 5; then if=0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62; then if=0(a,I,J) = Macro i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62; then if=0(a,I,J) = Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:62; then ProgramPart Relocated(J9,2) c= if=0(a,I,J) by A19,Lm5; then ProgramPart Relocated(J9,2) c= s3 by A21,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(J9,2) c= s5 by AMI_5:64; then A22: ProgramPart Relocated(J9,2) c= s5 by AMI_5:72; J0 c= J9 by SCMFSA6A:55; then ProgramPart Relocated(J0,2) c= ProgramPart Relocated(J9,2) by Th12; then A23: ProgramPart Relocated(J0,2) c= s5 by A22,XBOOLE_1:1; A24: IC s5 = s5.IC SCM+FSA by AMI_1:def 15 .= insloc 2 by A18,SCMFSA_2:95; s00,s3 equal_outside A by SCMFSA8A:14; then A25: s00 | D = s3 | D by SCMFSA6A:39; A26: now let a be Int-Location; thus s00.a = s3.a by A25,SCMFSA6A:38 .= s4.a by A13,SCMFSA_2:96 .= s5.a by A18,SCMFSA_2:95; end; now let f be FinSeq-Location; thus s00.f = s3.f by A25,SCMFSA6A:38 .= s4.f by A13,SCMFSA_2:96 .= s5.f by A18,SCMFSA_2:95; end; then A27: s00 | D = s5 | D by A26,SCMFSA6A:38; if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A28: if=0(a,I,J) c= s3 by SCMFSA6B:5; A29: 0 < card I + 2 by NAT_1:19; now thus card if=0(a,I,J) = card I + card J + (2 + 2) by SCMFSA8B:14 .= card I + card J + 2 + 2 by XCMPLX_1:1 .= card I + (card J + 2) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 2) by XCMPLX_1:1; end; then card J + 2 + 0 < card if=0(a,I,J) by A29,REAL_1:67; then A30: insloc (card J + 2) in dom if=0(a,I,J) by SCMFSA6A:15; set s9 = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2); A31: now thus IC s9 = IC (Computation s5).pseudo-LifeSpan(s00,J0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,J0) + 2 by A6,A7,A23,A24,A27,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,J0) + 2 by A4,Th50 .= insloc card ProgramPart J0 + 2 by A4,SCMFSA8A:def 5 .= insloc card J0 + 2 by AMI_5:72 .= insloc (card J0 + 2) by SCMFSA_4:def 1 .= insloc (card J + 2) by SCMFSA8A:34; end; then A32: CurInstr s9 = s9.insloc (card J + 2) by AMI_1:def 17 .= s3.insloc (card J + 2) by AMI_1:54 .= if=0(a,I,J).insloc (card J + 2) by A28,A30,GRFUNC_1:8 .= goto insloc (card I + card J + 3) by Th63; if=0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A33: if=0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if=0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:14 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if=0(a,I,J) by NAT_1:38; then A34: insloc (card I + card J + 3) in dom if=0(a,I,J) by SCMFSA6A:15; now thus IC (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = IC Following s9 by AMI_1:def 19 .= IC Exec(goto insloc (card I + card J + 3),s9) by A32,AMI_1:def 18 .= Exec(goto insloc (card I + card J + 3),s9).IC SCM+FSA by AMI_1:def 15 .= insloc (card I + card J + 3) by SCMFSA_2:95; end; then A35: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if=0(a,I,J).insloc (card I + card J + 3) by A33,A34,GRFUNC_1:8 .= halt SCM+FSA by Th61; then A36: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + (2 + 1)) = halt SCM+FSA by XCMPLX_1:1; A37: s3 is halting by A35,AMI_1:def 20; now let k be Nat; assume A38: CurInstr (Computation s3).k = halt SCM+FSA; assume A39: not pseudo-LifeSpan(s00,J0) + (1 + 2) <= k; CurInstr (Computation s3).0 = i by A12,AMI_1:def 19; then A40: k <> 0 & k <> 1 by A17,A38,SCMFSA_2:47,48,124; 2 < 3 & 0 <= card I + card J by NAT_1:18; then 0 + 2 < card I + card J + 3 by REAL_1:67; then A41: insloc 2 in dom if=0(a,I,J) & if=0(a,I,J).insloc 2 <> halt SCM+FSA by Th56; CurInstr (Computation s3).2 = (Computation s3).2.insloc 2 by A24,AMI_1:def 17 .= s3.insloc 2 by AMI_1:54 .= if=0(a,I,J).insloc 2 by A41,Th26; then 2 < k by A38,A40,A41,CQC_THE1:3; then consider k2 being Nat such that A42: 2 + k2 = k by NAT_1:28; k < pseudo-LifeSpan(s00,J0) + 1 + 2 by A39,XCMPLX_1:1; then k2 < pseudo-LifeSpan(s00,J0) + 1 by A42,AXIOMS:24; then A43: k2 <= pseudo-LifeSpan(s00,J0) by NAT_1:38; consider n being Nat such that A44: IC (Computation s00).k2 = insloc n by SCMFSA_2:21; A45: IC (Computation s3).k = IC (Computation s5).k2 by A42,AMI_1:51 .= IC (Computation s00).k2 + 2 by A6,A7,A23,A24,A27,A43,Th51 .= insloc (n + 2) by A44,SCMFSA_4:def 1; A46: k2 <= pseudo-LifeSpan(s,J0) by A4,A43,Th50; A47: now per cases by A46,REAL_1:def 5; suppose k2 = pseudo-LifeSpan(s,J0); then IC (Computation s00).k2 = insloc card ProgramPart J0 by A4,SCMFSA8A:def 5 .= insloc card J0 by AMI_5:72; then A48: n = card J0 by A44,SCMFSA_2:18 .= card J by SCMFSA8A:34; card I + card J + (2 + 1) = card I + (card J + (2 + 1)) by XCMPLX_1: 1 .= card J + 2 + 1 + card I by XCMPLX_1:1; then card J + 2 + 1 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A48,NAT_1:38; suppose k2 < pseudo-LifeSpan(s,J0); then insloc n in dom J0 by A4,A44,SCMFSA8A:31; then n < card J0 by SCMFSA6A:15; then n + 2 < card J0 + 2 by REAL_1:53; then A49: n + 2 < card J + 2 by SCMFSA8A:34; card I + card J + (1 + 2) = card I + card J + 1 + 2 by XCMPLX_1:1 .= card J + (card I + 1) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 1) by XCMPLX_1:1; then card J + 2 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A49,AXIOMS:22; end; then A50: insloc (n + 2) in dom if=0(a,I,J) & if=0(a,I,J).insloc (n + 2) <> halt SCM+FSA by Th56; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if=0(a,I,J).IC (Computation s3).k by A45,A50,Th26; hence contradiction by A38,A45,A47,Th56; end; then A51: LifeSpan s3 = pseudo-LifeSpan(s00,J0) + (2 + 1) by A36,A37,SCM_1:def 2; set s1 = s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0); s +* Initialized if=0(a,I,J) = Initialize s +* (if=0(a,I,J) +* Start-At insloc 0) & s +* Initialized (J ';' SCM+FSA-Stop) = Initialize s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0) by SCMFSA8A:13; then A52: s +* Initialized if=0(a,I,J) = s3 & s +* Initialized (J ';' SCM+FSA-Stop) = s1 by Th15; A53: J ';' SCM+FSA-Stop is_halting_on s & LifeSpan s1 = pseudo-LifeSpan(s,J0) by A4,Th58; then A54: s1 is halting by SCMFSA7B:def 8; A55: Directed J0 = J0 by SCMFSA8A:40; J0 ';' SCM+FSA-Stop = J ';' SCM+FSA-Stop & Directed J0 is_pseudo-closed_on s by A4,SCMFSA8A:40,41; then A56: (Computation s00).pseudo-LifeSpan(s,J0) | D = (Computation s1).(LifeSpan s1) | D by A53,A55,Th58; A57: LifeSpan s3 = pseudo-LifeSpan(s00,J0) + 2 + 1 by A51,XCMPLX_1:1; A58: s9 | D = (Computation s5).pseudo-LifeSpan(s00,J0) | D by AMI_1:51 .= (Computation s00).pseudo-LifeSpan(s00,J0) | D by A6,A7,A23,A24,A27,Th51; CurInstr s9 = s9.IC s9 by AMI_1:def 17 .= s3.insloc (card J + 2) by A31,AMI_1:54 .= if=0(a,I,J).insloc (card J + 2) by A30,Th26 .= goto insloc (card I + card J + 3) by Th63; then InsCode CurInstr s9 = 6 by SCMFSA_2:47; then InsCode CurInstr s9 in {0,6,7,8} by ENUMSET1:19; then A59: s9 | D = Exec(CurInstr s9,s9) | D by Th32 .= (Following s9) | D by AMI_1:def 18; thus IExec(if=0(a,I,J),ss) | D = IExec(if=0(a,I,J),s) | D by Th17 .= (Result (s +* Initialized if=0(a,I,J)) +* s | A) | D by SCMFSA6B:def 1 .= (Result s3) | D by A52,Th35 .= (Computation s3).(LifeSpan s3) | D by A37,SCMFSA6B:16 .= (Following s9) | D by A57,AMI_1:def 19 .= (Computation s00).pseudo-LifeSpan(s,J0) | D by A4,A58,A59,Th50 .= (Result s1) | D by A54,A56,SCMFSA6B:16 .= (Result (s +* Initialized (J ';' SCM+FSA-Stop)) +* s | A) | D by A52,Th35 .= IExec(J ';' SCM+FSA-Stop,s) | D by SCMFSA6B:def 1 .= IExec(J ';' SCM+FSA-Stop,ss) | D by Th17; end; theorem Th72: ::TMP171' for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.a <= 0 & Directed J is_pseudo-closed_on s holds if>0(a,I,J) is_halting_on s & if>0(a,I,J) is_closed_on s & LifeSpan (s +* (if>0(a,I,J) +* Start-At insloc 0)) = LifeSpan (s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0)) + 3 proof let s be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set J0 = Directed J; set s0 = Initialize s; set J9 = J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)); set s00 = s +* (J0 +* Start-At insloc 0); set s3 = s +* (if>0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set s5 = (Computation s3).2; set i = a >0_goto insloc (card J + 3); assume s.a <= 0; then A1: s0.a <= 0 by SCMFSA6C:3; A2: s3.a = s.a by Th28 .= s0.a by SCMFSA6C:3; assume A3: J0 is_pseudo-closed_on s; A4: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; A5: J0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A6: J0 is_pseudo-closed_on s00 by A3,Th52; A7: insloc 0 in dom if>0(a,I,J) by Th54; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A8: dom if>0(a,I,J) c= dom (if>0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A9: IC SCM+FSA in dom (if>0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A10: IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if>0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A9,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; s3.insloc 0 = (if>0(a,I,J) +* Start-At insloc 0).insloc 0 by A7,A8,FUNCT_4 :14 .= if>0(a,I,J).insloc 0 by A7,SCMFSA6B:7 .= i by Th55; then A11: CurInstr s3 = i by A10,AMI_1:def 17; A12: now thus (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; end; A13: now thus IC s4 = s4.IC SCM+FSA by AMI_1:def 15 .= Next IC s3 by A1,A2,A12,SCMFSA_2:97 .= insloc (0 + 1) by A10,SCMFSA_2:32; end; A14: insloc 1 in dom if>0(a,I,J) by Th54; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A15: dom if>0(a,I,J) c= dom (if>0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; now thus s4.insloc 1 = s3.insloc 1 by AMI_1:54 .= (if>0(a,I,J) +* Start-At insloc 0).insloc 1 by A14,A15,FUNCT_4:14 .= if>0(a,I,J).insloc 1 by A14,SCMFSA6B:7 .= goto insloc 2 by Th55; end; then A16: CurInstr s4 = goto insloc 2 by A13,AMI_1:def 17; A17: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(goto insloc 2,s4) by A16,AMI_1:def 18; A18: card Macro i = 2 by SCMFSA7B:6; A19: if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A20: if>0(a,I,J) c= s3 by A19,XBOOLE_1:1; if>0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by A4,SCMFSA6A:def 5; then if>0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62; then if>0(a,I,J) = Macro i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62; then if>0(a,I,J) = Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:62; then ProgramPart Relocated(J9,2) c= if>0(a,I,J) by A18,Lm5; then ProgramPart Relocated(J9,2) c= s3 by A20,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(J9,2) c= s5 by AMI_5:64; then A21: ProgramPart Relocated(J9,2) c= s5 by AMI_5:72; J0 c= J9 by SCMFSA6A:55; then ProgramPart Relocated(J0,2) c= ProgramPart Relocated(J9,2) by Th12; then A22: ProgramPart Relocated(J0,2) c= s5 by A21,XBOOLE_1:1; A23: IC s5 = s5.IC SCM+FSA by AMI_1:def 15 .= insloc 2 by A17,SCMFSA_2:95; s00,s3 equal_outside A by SCMFSA8A:14; then A24: s00 | D = s3 | D by SCMFSA6A:39; A25: now let a be Int-Location; thus s00.a = s3.a by A24,SCMFSA6A:38 .= s4.a by A12,SCMFSA_2:97 .= s5.a by A17,SCMFSA_2:95; end; now let f be FinSeq-Location; thus s00.f = s3.f by A24,SCMFSA6A:38 .= s4.f by A12,SCMFSA_2:97 .= s5.f by A17,SCMFSA_2:95; end; then A26: s00 | D = s5 | D by A25,SCMFSA6A:38; if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A27: if>0(a,I,J) c= s3 by SCMFSA6B:5; A28: 0 < card I + 2 by NAT_1:19; now thus card if>0(a,I,J) = card I + card J + (2 + 2) by SCMFSA8B:15 .= card I + card J + 2 + 2 by XCMPLX_1:1 .= card I + (card J + 2) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 2) by XCMPLX_1:1; end; then A29: card J + 2 + 0 < card if>0(a,I,J) by A28,REAL_1:67; then A30: insloc (card J + 2) in dom if>0(a,I,J) by SCMFSA6A:15; set ss = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2); A31: now thus IC ss = IC (Computation s5).pseudo-LifeSpan(s00,J0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,J0) + 2 by A5,A6,A22,A23,A26,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,J0) + 2 by A3,Th50 .= insloc card ProgramPart J0 + 2 by A3,SCMFSA8A:def 5 .= insloc card J0 + 2 by AMI_5:72 .= insloc (card J0 + 2) by SCMFSA_4:def 1 .= insloc (card J + 2) by SCMFSA8A:34; end; then A32: CurInstr ss = ss.insloc (card J + 2) by AMI_1:def 17 .= s3.insloc (card J + 2) by AMI_1:54 .= if>0(a,I,J).insloc (card J + 2) by A27,A30,GRFUNC_1:8 .= goto insloc (card I + card J + 3) by Th64; if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A33: if>0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if>0(a,I,J) by NAT_1:38; then A34: insloc (card I + card J + 3) in dom if>0(a,I,J) by SCMFSA6A:15; A35: now thus IC (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = IC Following ss by AMI_1:def 19 .= IC Exec(goto insloc (card I + card J + 3),ss) by A32,AMI_1:def 18 .= Exec(goto insloc (card I + card J + 3),ss).IC SCM+FSA by AMI_1:def 15 .= insloc (card I + card J + 3) by SCMFSA_2:95; end; then A36: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if>0(a,I,J).insloc (card I + card J + 3) by A33,A34,GRFUNC_1:8 .= halt SCM+FSA by Th62; then A37: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + (2 + 1)) = halt SCM+FSA by XCMPLX_1:1; A38: s3 is halting by A36,AMI_1:def 20; hence if>0(a,I,J) is_halting_on s by SCMFSA7B:def 8; now let k be Nat; per cases by NAT_1:19; suppose k = 0; then (Computation s3).k = s3 by AMI_1:def 19; then IC (Computation s3).k = insloc 0 by Th31; hence IC (Computation s3).k in dom if>0(a,I,J) by Th54; suppose 0 < k & k = 1; hence IC (Computation s3).k in dom if>0(a,I,J) by A13,Th54; suppose A39: 0 < k & k <> 1 & k < pseudo-LifeSpan(s00,J0) + 2; then 0 + 1 <= k by INT_1:20; then 1 < k by A39,REAL_1:def 5; then 0 + (1 + 1) <= k by INT_1:20; then consider k2 being Nat such that A40: 2 + k2 = k by NAT_1:28; A41: k2 < pseudo-LifeSpan(s00,J0) by A39,A40,AXIOMS:24; then A42: k2 < pseudo-LifeSpan(s,J0) by A3,Th50; A43: IC (Computation s3).k = IC (Computation s5).k2 by A40,AMI_1:51 .= IC (Computation s00).k2 + 2 by A5,A6,A22,A23,A26,A41,Th51; consider n being Nat such that A44: IC (Computation s00).k2 = insloc n by SCMFSA_2:21; insloc n in dom J0 by A3,A42,A44,SCMFSA8A:31; then n < card J0 by SCMFSA6A:15; then n + 2 < card J0 + 2 by REAL_1:53; then A45: n + 2 < card J + 2 by SCMFSA8A:34; card if>0(a,I,J) = card I + card J + (2 + 2) by SCMFSA8B:15 .= card I + card J + 2 + 2 by XCMPLX_1:1 .= card J + (card I + 2) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 2) by XCMPLX_1:1; then card J + 2 <= card if>0(a,I,J) by NAT_1:37; then n + 2 < card if>0(a,I,J) by A45,AXIOMS:22; then insloc (n + 2) in dom if>0(a,I,J) by SCMFSA6A:15; hence IC (Computation s3).k in dom if>0(a,I,J) by A43,A44,SCMFSA_4:def 1; suppose A46: 0 < k & k <> 1 & pseudo-LifeSpan(s00,J0) + 2 <= k; hereby per cases by A46,REAL_1:def 5; suppose pseudo-LifeSpan(s00,J0) + 2 = k; hence IC (Computation s3).k in dom if>0(a,I,J) by A29,A31,SCMFSA6A:15; suppose pseudo-LifeSpan(s00,J0) + 2 < k; then pseudo-LifeSpan(s00,J0) + 2 + 1 <= k by INT_1:20; hence IC (Computation s3).k in dom if>0(a,I,J) by A34,A35,A36,AMI_1:52 ; end; end; hence if>0(a,I,J) is_closed_on s by SCMFSA7B:def 7; now let k be Nat; assume A47: CurInstr (Computation s3).k = halt SCM+FSA; assume A48: not pseudo-LifeSpan(s00,J0) + (1 + 2) <= k; CurInstr (Computation s3).0 = i by A11,AMI_1:def 19; then A49: k <> 0 & k <> 1 by A16,A47,SCMFSA_2:47,49,124; 2 < 3 & 0 <= card I + card J by NAT_1:18; then 0 + 2 < card I + card J + 3 by REAL_1:67; then A50: insloc 2 in dom if>0(a,I,J) & if>0(a,I,J).insloc 2 <> halt SCM+FSA by Th57; CurInstr (Computation s3).2 = (Computation s3).2.insloc 2 by A23,AMI_1:def 17 .= s3.insloc 2 by AMI_1:54 .= if>0(a,I,J).insloc 2 by A50,Th26; then 2 < k by A47,A49,A50,CQC_THE1:3; then consider k2 being Nat such that A51: 2 + k2 = k by NAT_1:28; k < pseudo-LifeSpan(s00,J0) + 1 + 2 by A48,XCMPLX_1:1; then k2 < pseudo-LifeSpan(s00,J0) + 1 by A51,AXIOMS:24; then A52: k2 <= pseudo-LifeSpan(s00,J0) by NAT_1:38; consider n being Nat such that A53: IC (Computation s00).k2 = insloc n by SCMFSA_2:21; A54: IC (Computation s3).k = IC (Computation s5).k2 by A51,AMI_1:51 .= IC (Computation s00).k2 + 2 by A5,A6,A22,A23,A26,A52,Th51 .= insloc (n + 2) by A53,SCMFSA_4:def 1; A55: k2 <= pseudo-LifeSpan(s,J0) by A3,A52,Th50; A56: now per cases by A55,REAL_1:def 5; suppose k2 = pseudo-LifeSpan(s,J0); then IC (Computation s00).k2 = insloc card ProgramPart J0 by A3,SCMFSA8A:def 5 .= insloc card J0 by AMI_5:72; then A57: n = card J0 by A53,SCMFSA_2:18 .= card J by SCMFSA8A:34; card I + card J + (2 + 1) = card I + (card J + (2 + 1)) by XCMPLX_1: 1 .= card J + 2 + 1 + card I by XCMPLX_1:1; then card J + 2 + 1 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A57,NAT_1:38; suppose k2 < pseudo-LifeSpan(s,J0); then insloc n in dom J0 by A3,A53,SCMFSA8A:31; then n < card J0 by SCMFSA6A:15; then n + 2 < card J0 + 2 by REAL_1:53; then A58: n + 2 < card J + 2 by SCMFSA8A:34; card I + card J + (1 + 2) = card I + card J + 1 + 2 by XCMPLX_1:1 .= card J + (card I + 1) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 1) by XCMPLX_1:1; then card J + 2 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A58,AXIOMS:22; end; then A59: insloc (n + 2) in dom if>0(a,I,J) & if>0(a,I,J).insloc (n + 2) <> halt SCM+FSA by Th57; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if>0(a,I,J).IC (Computation s3).k by A54,A59,Th26; hence contradiction by A47,A54,A56,Th57; end; then A60: LifeSpan s3 = pseudo-LifeSpan(s00,J0) + 3 by A37,A38,SCM_1:def 2; pseudo-LifeSpan(s,J0) = LifeSpan (s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0)) by A3,Th58; hence LifeSpan (s +* (if>0(a,I,J) +* Start-At insloc 0)) = LifeSpan (s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0)) + 3 by A3,A60,Th50; end; theorem Th73: ::TMP17' for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st s.intloc 0 = 1 & s.a <= 0 & Directed J is_pseudo-closed_on s holds IExec(if>0(a,I,J),s) | (Int-Locations \/ FinSeq-Locations) = IExec(J ';' SCM+FSA-Stop,s) | (Int-Locations \/ FinSeq-Locations) proof let ss be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; set J0 = Directed J; set s = Initialize ss; set s0 = Initialize s; set J9 = J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)); set s00 = s +* (J0 +* Start-At insloc 0); set s3 = s +* (if>0(a,I,J) +* Start-At insloc 0); set s4 = (Computation s3).1; set s5 = (Computation s3).2; set i = a >0_goto insloc (card J + 3); assume A1: ss.intloc 0 = 1; assume ss.a <= 0; then s.a <= 0 by SCMFSA6C:3; then A2: s0.a <= 0 by SCMFSA6C:3; A3: s3.a = s.a by Th28 .= s0.a by SCMFSA6C:3; assume J0 is_pseudo-closed_on ss; then A4: J0 is_pseudo-closed_on s by A1,Th53; A5: if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; A6: J0 +* Start-At insloc 0 c= s00 by FUNCT_4:26; s | D = s00 | D by SCMFSA8A:11; then A7: J0 is_pseudo-closed_on s00 by A4,Th52; A8: insloc 0 in dom if>0(a,I,J) by Th54; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A9: dom if>0(a,I,J) c= dom (if>0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; A10: IC SCM+FSA in dom (if>0(a,I,J) +* Start-At insloc 0) by SF_MASTR:65; A11: IC s3 = s3.IC SCM+FSA by AMI_1:def 15 .= (if>0(a,I,J) +* Start-At insloc 0).IC SCM+FSA by A10,FUNCT_4:14 .= insloc 0 by SF_MASTR:66; s3.insloc 0 = (if>0(a,I,J) +* Start-At insloc 0).insloc 0 by A8,A9,FUNCT_4 :14 .= if>0(a,I,J).insloc 0 by A8,SCMFSA6B:7 .= i by Th55; then A12: CurInstr s3 = i by A11,AMI_1:def 17; A13: now thus (Computation s3).(0 + 1) = Following (Computation s3).0 by AMI_1:def 19 .= Following s3 by AMI_1:def 19 .= Exec(i,s3) by A12,AMI_1:def 18; end; A14: now thus IC s4 = s4.IC SCM+FSA by AMI_1:def 15 .= Next IC s3 by A2,A3,A13,SCMFSA_2:97 .= insloc (0 + 1) by A11,SCMFSA_2:32; end; A15: insloc 1 in dom if>0(a,I,J) by Th54; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A16: dom if>0(a,I,J) c= dom (if>0(a,I,J) +* Start-At insloc 0) by GRFUNC_1:8; now thus s4.insloc 1 = s3.insloc 1 by AMI_1:54 .= (if>0(a,I,J) +* Start-At insloc 0).insloc 1 by A15,A16,FUNCT_4:14 .= if>0(a,I,J).insloc 1 by A15,SCMFSA6B:7 .= goto insloc 2 by Th55; end; then A17: CurInstr s4 = goto insloc 2 by A14,AMI_1:def 17; A18: (Computation s3).(1 + 1) = Following s4 by AMI_1:def 19 .= Exec(goto insloc 2,s4) by A17,AMI_1:def 18; A19: card Macro i = 2 by SCMFSA7B:6; A20: if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; if>0(a,I,J) c= if>0(a,I,J) +* Start-At insloc 0 by SCMFSA8A:9; then A21: if>0(a,I,J) c= s3 by A20,XBOOLE_1:1; if>0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by A5,SCMFSA6A:def 5; then if>0(a,I,J) = Macro i ';' J ';' Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop) by SCMFSA6A:62; then if>0(a,I,J) = Macro i ';' J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop)) by SCMFSA6A:62; then if>0(a,I,J) = Macro i ';' (J ';' (Goto insloc (card I + 1) ';' (I ';' SCM+FSA-Stop))) by SCMFSA6A:62; then ProgramPart Relocated(J9,2) c= if>0(a,I,J) by A19,Lm5; then ProgramPart Relocated(J9,2) c= s3 by A21,XBOOLE_1:1; then ProgramPart ProgramPart Relocated(J9,2) c= s5 by AMI_5:64; then A22: ProgramPart Relocated(J9,2) c= s5 by AMI_5:72; J0 c= J9 by SCMFSA6A:55; then ProgramPart Relocated(J0,2) c= ProgramPart Relocated(J9,2) by Th12; then A23: ProgramPart Relocated(J0,2) c= s5 by A22,XBOOLE_1:1; A24: IC s5 = s5.IC SCM+FSA by AMI_1:def 15 .= insloc 2 by A18,SCMFSA_2:95; s00,s3 equal_outside A by SCMFSA8A:14; then A25: s00 | D = s3 | D by SCMFSA6A:39; A26: now let a be Int-Location; thus s00.a = s3.a by A25,SCMFSA6A:38 .= s4.a by A13,SCMFSA_2:97 .= s5.a by A18,SCMFSA_2:95; end; now let f be FinSeq-Location; thus s00.f = s3.f by A25,SCMFSA6A:38 .= s4.f by A13,SCMFSA_2:97 .= s5.f by A18,SCMFSA_2:95; end; then A27: s00 | D = s5 | D by A26,SCMFSA6A:38; if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A28: if>0(a,I,J) c= s3 by SCMFSA6B:5; A29: 0 < card I + 2 by NAT_1:19; now thus card if>0(a,I,J) = card I + card J + (2 + 2) by SCMFSA8B:15 .= card I + card J + 2 + 2 by XCMPLX_1:1 .= card I + (card J + 2) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 2) by XCMPLX_1:1; end; then card J + 2 + 0 < card if>0(a,I,J) by A29,REAL_1:67; then A30: insloc (card J + 2) in dom if>0(a,I,J) by SCMFSA6A:15; set s9 = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2); A31: now thus IC s9 = IC (Computation s5).pseudo-LifeSpan(s00,J0) by AMI_1:51 .= IC (Computation s00).pseudo-LifeSpan(s00,J0) + 2 by A6,A7,A23,A24,A27,Th51 .= IC (Computation s00).pseudo-LifeSpan(s,J0) + 2 by A4,Th50 .= insloc card ProgramPart J0 + 2 by A4,SCMFSA8A:def 5 .= insloc card J0 + 2 by AMI_5:72 .= insloc (card J0 + 2) by SCMFSA_4:def 1 .= insloc (card J + 2) by SCMFSA8A:34; end; then A32: CurInstr s9 = s9.insloc (card J + 2) by AMI_1:def 17 .= s3.insloc (card J + 2) by AMI_1:54 .= if>0(a,I,J).insloc (card J + 2) by A28,A30,GRFUNC_1:8 .= goto insloc (card I + card J + 3) by Th64; if>0(a,I,J) +* Start-At insloc 0 c= s3 by FUNCT_4:26; then A33: if>0(a,I,J) c= s3 by SCMFSA6B:5; now thus card if>0(a,I,J) = card I + card J + (3 + 1) by SCMFSA8B:15 .= card I + card J + 3 + 1 by XCMPLX_1:1; end; then card I + card J + 3 < card if>0(a,I,J) by NAT_1:38; then A34: insloc (card I + card J + 3) in dom if>0(a,I,J) by SCMFSA6A:15; now thus IC (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = IC Following s9 by AMI_1:def 19 .= IC Exec(goto insloc (card I + card J + 3),s9) by A32,AMI_1:def 18 .= Exec(goto insloc (card I + card J + 3),s9).IC SCM+FSA by AMI_1:def 15 .= insloc (card I + card J + 3) by SCMFSA_2:95; end; then A35: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1) = (Computation s3).(pseudo-LifeSpan(s00,J0) + 2 + 1). insloc (card I + card J + 3) by AMI_1:def 17 .= s3.insloc (card I + card J + 3) by AMI_1:54 .= if>0(a,I,J).insloc (card I + card J + 3) by A33,A34,GRFUNC_1:8 .= halt SCM+FSA by Th62; then A36: CurInstr (Computation s3).(pseudo-LifeSpan(s00,J0) + (2 + 1)) = halt SCM+FSA by XCMPLX_1:1; A37: s3 is halting by A35,AMI_1:def 20; now let k be Nat; assume A38: CurInstr (Computation s3).k = halt SCM+FSA; assume A39: not pseudo-LifeSpan(s00,J0) + (1 + 2) <= k; CurInstr (Computation s3).0 = i by A12,AMI_1:def 19; then A40: k <> 0 & k <> 1 by A17,A38,SCMFSA_2:47,49,124; 2 < 3 & 0 <= card I + card J by NAT_1:18; then 0 + 2 < card I + card J + 3 by REAL_1:67; then A41: insloc 2 in dom if>0(a,I,J) & if>0(a,I,J).insloc 2 <> halt SCM+FSA by Th57; CurInstr (Computation s3).2 = (Computation s3).2.insloc 2 by A24,AMI_1:def 17 .= s3.insloc 2 by AMI_1:54 .= if>0(a,I,J).insloc 2 by A41,Th26; then 2 < k by A38,A40,A41,CQC_THE1:3; then consider k2 being Nat such that A42: 2 + k2 = k by NAT_1:28; k < pseudo-LifeSpan(s00,J0) + 1 + 2 by A39,XCMPLX_1:1; then k2 < pseudo-LifeSpan(s00,J0) + 1 by A42,AXIOMS:24; then A43: k2 <= pseudo-LifeSpan(s00,J0) by NAT_1:38; consider n being Nat such that A44: IC (Computation s00).k2 = insloc n by SCMFSA_2:21; A45: IC (Computation s3).k = IC (Computation s5).k2 by A42,AMI_1:51 .= IC (Computation s00).k2 + 2 by A6,A7,A23,A24,A27,A43,Th51 .= insloc (n + 2) by A44,SCMFSA_4:def 1; A46: k2 <= pseudo-LifeSpan(s,J0) by A4,A43,Th50; A47: now per cases by A46,REAL_1:def 5; suppose k2 = pseudo-LifeSpan(s,J0); then IC (Computation s00).k2 = insloc card ProgramPart J0 by A4,SCMFSA8A:def 5 .= insloc card J0 by AMI_5:72; then A48: n = card J0 by A44,SCMFSA_2:18 .= card J by SCMFSA8A:34; card I + card J + (2 + 1) = card I + (card J + (2 + 1)) by XCMPLX_1: 1 .= card J + 2 + 1 + card I by XCMPLX_1:1; then card J + 2 + 1 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A48,NAT_1:38; suppose k2 < pseudo-LifeSpan(s,J0); then insloc n in dom J0 by A4,A44,SCMFSA8A:31; then n < card J0 by SCMFSA6A:15; then n + 2 < card J0 + 2 by REAL_1:53; then A49: n + 2 < card J + 2 by SCMFSA8A:34; card I + card J + (1 + 2) = card I + card J + 1 + 2 by XCMPLX_1:1 .= card J + (card I + 1) + 2 by XCMPLX_1:1 .= card J + 2 + (card I + 1) by XCMPLX_1:1; then card J + 2 <= card I + card J + 3 by NAT_1:29; hence n + 2 < card I + card J + 3 by A49,AXIOMS:22; end; then A50: insloc (n + 2) in dom if>0(a,I,J) & if>0(a,I,J).insloc (n + 2) <> halt SCM+FSA by Th57; CurInstr (Computation s3).k = (Computation s3).k.IC (Computation s3).k by AMI_1:def 17 .= s3.IC (Computation s3).k by AMI_1:54 .= if>0(a,I,J).IC (Computation s3).k by A45,A50,Th26; hence contradiction by A38,A45,A47,Th57; end; then A51: LifeSpan s3 = pseudo-LifeSpan(s00,J0) + (2 + 1) by A36,A37,SCM_1:def 2; set s1 = s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0); s +* Initialized if>0(a,I,J) = Initialize s +* (if>0(a,I,J) +* Start-At insloc 0) & s +* Initialized (J ';' SCM+FSA-Stop) = Initialize s +* (J ';' SCM+FSA-Stop +* Start-At insloc 0) by SCMFSA8A:13; then A52: s +* Initialized if>0(a,I,J) = s3 & s +* Initialized (J ';' SCM+FSA-Stop) = s1 by Th15; A53: J ';' SCM+FSA-Stop is_halting_on s & LifeSpan s1 = pseudo-LifeSpan(s,J0) by A4,Th58; then A54: s1 is halting by SCMFSA7B:def 8; A55: Directed J0 = J0 by SCMFSA8A:40; J0 ';' SCM+FSA-Stop = J ';' SCM+FSA-Stop & Directed J0 is_pseudo-closed_on s by A4,SCMFSA8A:40,41; then A56: (Computation s00).pseudo-LifeSpan(s,J0) | D = (Computation s1).(LifeSpan s1) | D by A53,A55,Th58; A57: LifeSpan s3 = pseudo-LifeSpan(s00,J0) + 2 + 1 by A51,XCMPLX_1:1; A58: s9 | D = (Computation s5).pseudo-LifeSpan(s00,J0) | D by AMI_1:51 .= (Computation s00).pseudo-LifeSpan(s00,J0) | D by A6,A7,A23,A24,A27,Th51; CurInstr s9 = s9.IC s9 by AMI_1:def 17 .= s3.insloc (card J + 2) by A31,AMI_1:54 .= if>0(a,I,J).insloc (card J + 2) by A30,Th26 .= goto insloc (card I + card J + 3) by Th64; then InsCode CurInstr s9 = 6 by SCMFSA_2:47; then InsCode CurInstr s9 in {0,6,7,8} by ENUMSET1:19; then A59: s9 | D = Exec(CurInstr s9,s9) | D by Th32 .= (Following s9) | D by AMI_1:def 18; thus IExec(if>0(a,I,J),ss) | D = IExec(if>0(a,I,J),s) | D by Th17 .= (Result (s +* Initialized if>0(a,I,J)) +* s | A) | D by SCMFSA6B:def 1 .= (Result s3) | D by A52,Th35 .= (Computation s3).(LifeSpan s3) | D by A37,SCMFSA6B:16 .= (Following s9) | D by A57,AMI_1:def 19 .= (Computation s00).pseudo-LifeSpan(s,J0) | D by A4,A58,A59,Th50 .= (Result s1) | D by A54,A56,SCMFSA6B:16 .= (Result (s +* Initialized (J ';' SCM+FSA-Stop)) +* s | A) | D by A52,Th35 .= IExec(J ';' SCM+FSA-Stop,s) | D by SCMFSA6B:def 1 .= IExec(J ';' SCM+FSA-Stop,ss) | D by Th17; end; theorem for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st Directed I is_pseudo-closed_on s & Directed J is_pseudo-closed_on s holds if=0(a,I,J) is_closed_on s & if=0(a,I,J) is_halting_on s proof let s be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; assume A1: Directed I is_pseudo-closed_on s; assume A2: Directed J is_pseudo-closed_on s; hereby per cases; suppose A3: s.a = 0; hence if=0(a,I,J) is_closed_on s by A1,Th66; thus if=0(a,I,J) is_halting_on s by A1,A3,Th66; suppose A4: s.a <> 0; hence if=0(a,I,J) is_closed_on s by A2,Th70; thus if=0(a,I,J) is_halting_on s by A2,A4,Th70; end; end; theorem for s being State of SCM+FSA, I,J being Macro-Instruction, a being read-write Int-Location st Directed I is_pseudo-closed_on s & Directed J is_pseudo-closed_on s holds if>0(a,I,J) is_closed_on s & if>0(a,I,J) is_halting_on s proof let s be State of SCM+FSA; let I,J be Macro-Instruction; let a be read-write Int-Location; assume A1: Directed I is_pseudo-closed_on s; assume A2: Directed J is_pseudo-closed_on s; hereby per cases; suppose A3: s.a > 0; hence if>0(a,I,J) is_closed_on s by A1,Th68; thus if>0(a,I,J) is_halting_on s by A1,A3,Th68; suppose A4: s.a <= 0; hence if>0(a,I,J) is_closed_on s by A2,Th72; thus if>0(a,I,J) is_halting_on s by A2,A4,Th72; end; end; theorem Th76: ::TG8' for I being Macro-Instruction, a being Int-Location holds I does_not_destroy a implies Directed I does_not_destroy a proof let I be Macro-Instruction; let a be Int-Location; assume A1: I does_not_destroy a; Directed I = Directed(I,insloc card I) by SCMFSA8A:18; hence thesis by A1,SCMFSA8A:27; end; theorem Th77: ::Td1(@BBB8) for i being Instruction of SCM+FSA, a being Int-Location holds i does_not_destroy a implies Macro i does_not_destroy a proof let i be Instruction of SCM+FSA; let a be Int-Location; assume A1: i does_not_destroy a; A2: rng Macro i = {i,halt SCM+FSA} by Th41; now let ii be Instruction of SCM+FSA; assume ii in rng Macro i; then ii = i or ii = halt SCM+FSA by A2,TARSKI:def 2; hence ii does_not_destroy a by A1,SCMFSA7B:11; end; hence Macro i does_not_destroy a by SCMFSA7B:def 4; end; theorem Th78: ::R0' for a being Int-Location holds halt SCM+FSA does_not_refer a proof let a be Int-Location; now let b be Int-Location; let l be Instruction-Location of SCM+FSA; let f be FinSeq-Location; thus b := a <> halt SCM+FSA by SCMFSA_2:42,124; thus AddTo(b,a) <> halt SCM+FSA by SCMFSA_2:43,124; thus SubFrom(b,a) <> halt SCM+FSA by SCMFSA_2:44,124; thus MultBy(b,a) <> halt SCM+FSA by SCMFSA_2:45,124; thus Divide(a,b) <> halt SCM+FSA & Divide(b,a) <> halt SCM+FSA by SCMFSA_2:46,124; thus a =0_goto l <> halt SCM+FSA by SCMFSA_2:48,124; thus a >0_goto l <> halt SCM+FSA by SCMFSA_2:49,124; thus b :=(f,a) <> halt SCM+FSA by SCMFSA_2:50,124; thus (f,b):= a <> halt SCM+FSA & (f,a):= b <> halt SCM+FSA by SCMFSA_2:51,124; thus f :=<0,...,0> a <> halt SCM+FSA by SCMFSA_2:53,124; end; hence halt SCM+FSA does_not_refer a by SCMFSA7B:def 1; end; theorem for a,b,c being Int-Location holds a <> b implies AddTo(c,b) does_not_refer a proof let a,b,c be Int-Location; assume A1: a <> b; now let e be Int-Location; let l be Instruction-Location of SCM+FSA; let f be FinSeq-Location; A2: 2 <> 1 & 2 <> 3 & 2 <> 4 & 2 <> 5 & 2 <> 7 & 2 <> 8 & 2 <> 9 & 2 <> 10 & 2 <> 12 & InsCode AddTo(c,b) = 2 by SCMFSA_2:43; hence e := a <> AddTo(c,b) by SCMFSA_2:42; thus AddTo(e,a) <> AddTo(c,b) by A1,SF_MASTR:6; thus SubFrom(e,a) <> AddTo(c,b) by A2,SCMFSA_2:44; thus MultBy(e,a) <> AddTo(c,b) by A2,SCMFSA_2:45; thus Divide(a,e) <> AddTo(c,b) & Divide(e,a) <> AddTo(c,b) by A2,SCMFSA_2:46; thus a =0_goto l <> AddTo(c,b) by A2,SCMFSA_2:48; thus a >0_goto l <> AddTo(c,b) by A2,SCMFSA_2:49; thus e :=(f,a) <> AddTo(c,b) by A2,SCMFSA_2:50; thus (f,e):= a <> AddTo(c,b) & (f,a):= e <> AddTo(c,b) by A2,SCMFSA_2:51 ; thus f :=<0,...,0> a <> AddTo(c,b) by A2,SCMFSA_2:53; end; hence AddTo(c,b) does_not_refer a by SCMFSA7B:def 1; end; theorem for i being Instruction of SCM+FSA, a being Int-Location holds i does_not_refer a implies Macro i does_not_refer a proof let i be Instruction of SCM+FSA; let a be Int-Location; assume A1: i does_not_refer a; A2: rng Macro i = {i,halt SCM+FSA} by Th41; now let ii be Instruction of SCM+FSA; assume ii in rng Macro i; then ii = i or ii = halt SCM+FSA by A2,TARSKI:def 2; hence ii does_not_refer a by A1,Th78; end; hence Macro i does_not_refer a by SCMFSA7B:def 2; end; theorem Th81: ::TG9' for I,J being Macro-Instruction, a being Int-Location holds I does_not_destroy a & J does_not_destroy a implies I ';' J does_not_destroy a proof let I,J be Macro-Instruction; let a be Int-Location; assume A1: I does_not_destroy a & J does_not_destroy a; A2: I ';' J = Directed I +* ProgramPart Relocated(J,card I) by SCMFSA6A:def 4; A3: Directed I does_not_destroy a by A1,Th76; ProgramPart Relocated(J,card I) does_not_destroy a by A1,SCMFSA8A:23; hence I ';' J does_not_destroy a by A2,A3,SCMFSA8A:25; end; theorem Th82: ::T281220 for J being Macro-Instruction, i being Instruction of SCM+FSA, a being Int-Location st i does_not_destroy a & J does_not_destroy a holds i ';' J does_not_destroy a proof let J be Macro-Instruction; let i be Instruction of SCM+FSA; let a be Int-Location; assume A1: i does_not_destroy a & J does_not_destroy a; then A2: Macro i does_not_destroy a by Th77; i ';' J = Macro i ';' J by SCMFSA6A:def 5; hence i ';' J does_not_destroy a by A1,A2,Th81; end; theorem for I being Macro-Instruction, j being Instruction of SCM+FSA, a being Int-Location st I does_not_destroy a & j does_not_destroy a holds I ';' j does_not_destroy a proof let I be Macro-Instruction; let j be Instruction of SCM+FSA; let a be Int-Location; assume A1: I does_not_destroy a & j does_not_destroy a; then A2: Macro j does_not_destroy a by Th77; I ';' j = I ';' Macro j by SCMFSA6A:def 6; hence I ';' j does_not_destroy a by A1,A2,Th81; end; theorem for i,j being Instruction of SCM+FSA, a being Int-Location st i does_not_destroy a & j does_not_destroy a holds i ';' j does_not_destroy a proof let i,j be Instruction of SCM+FSA; let a be Int-Location; assume i does_not_destroy a & j does_not_destroy a; then A1: Macro i does_not_destroy a & Macro j does_not_destroy a by Th77; i ';' j = Macro i ';' Macro j by SCMFSA6A:def 7; hence i ';' j does_not_destroy a by A1,Th81; end; theorem Th85: ::TQ7' for a being Int-Location holds SCM+FSA-Stop does_not_destroy a proof let a be Int-Location; now let i be Instruction of SCM+FSA; assume A1: i in rng SCM+FSA-Stop; rng SCM+FSA-Stop = {halt SCM+FSA} by CQC_LANG:5,SCMFSA_4:def 5; then i = halt SCM+FSA by A1,TARSKI:def 1; hence i does_not_destroy a by SCMFSA7B:11; end; hence thesis by SCMFSA7B:def 4; end; theorem Th86: ::T27749 for a being Int-Location, l being Instruction-Location of SCM+FSA holds Goto l does_not_destroy a proof let a be Int-Location; let l be Instruction-Location of SCM+FSA; now let i be Instruction of SCM+FSA; assume A1: i in rng Goto l; Goto l = insloc 0 .--> goto l by SCMFSA8A:def 2; then rng Goto l = {goto l} by CQC_LANG:5; then i = goto l by A1,TARSKI:def 1; hence i does_not_destroy a by SCMFSA7B:17; end; hence thesis by SCMFSA7B:def 4; end; theorem Th87: ::T200724' for s being State of SCM+FSA, I being Macro-Instruction st I is_halting_on Initialize s holds (for a being read-write Int-Location holds IExec(I,s).a = (Computation (Initialize s +* (I +* Start-At insloc 0))). (LifeSpan (Initialize s +* (I +* Start-At insloc 0))).a) & for f being FinSeq-Location holds IExec(I,s).f = (Computation (Initialize s +* (I +* Start-At insloc 0))). (LifeSpan (Initialize s +* (I +* Start-At insloc 0))).f proof let s be State of SCM+FSA; let I be Macro-Instruction; set s0 = Initialize s; set s1 = s0 +* (I +* Start-At insloc 0); assume I is_halting_on s0; then A1: s1 is halting by SCMFSA7B:def 8; hereby let a be read-write Int-Location; not a in A by SCMFSA_2:84; then not a in dom s /\ A by XBOOLE_0:def 3; then A2: a in dom Result s1 & not a in dom (s | A) by RELAT_1:90,SCMFSA_2:66; s +* Initialized I = s1 by SCMFSA8A:13; hence IExec(I,s).a = (Result s1 +* s | A).a by SCMFSA6B:def 1 .= (Result s1).a by A2,FUNCT_4:12 .= (Computation s1).(LifeSpan s1).a by A1,SCMFSA6B:16; end; let f be FinSeq-Location; not f in A by SCMFSA_2:85; then not f in dom s /\ A by XBOOLE_0:def 3; then A3: f in dom Result s1 & not f in dom (s | A) by RELAT_1:90,SCMFSA_2:67; s +* Initialized I = s1 by SCMFSA8A:13; hence IExec(I,s).f = (Result s1 +* s | A).f by SCMFSA6B:def 1 .= (Result s1).f by A3,FUNCT_4:12 .= (Computation s1).(LifeSpan s1).f by A1,SCMFSA6B:16; end; theorem Th88: ::T200724 for s being State of SCM+FSA, I being parahalting Macro-Instruction, a being read-write Int-Location holds IExec(I,s).a = (Computation (Initialize s +* (I +* Start-At insloc 0))). (LifeSpan (Initialize s +* (I +* Start-At insloc 0))).a proof let s be State of SCM+FSA; let I be parahalting Macro-Instruction; let a be read-write Int-Location; I is_halting_on Initialize s by SCMFSA7B:25; hence thesis by Th87; end; theorem Th89: ::TMP29' for s being State of SCM+FSA, I being Macro-Instruction, a being Int-Location,k being Nat st I is_closed_on Initialize s & I is_halting_on Initialize s & I does_not_destroy a holds IExec(I,s).a = (Computation (Initialize s +* (I +* Start-At insloc 0))).k.a proof let s be State of SCM+FSA; let I be Macro-Instruction; let a be Int-Location; let k be Nat; assume A1: I is_closed_on Initialize s; assume A2: I is_halting_on Initialize s; assume A3: I does_not_destroy a; set s0 = Initialize s; set s1 = s0 +* (I +* Start-At insloc 0); A4: s1 is halting by A2,SCMFSA7B:def 8; not a in A by SCMFSA_2:84; then not a in dom s /\ A by XBOOLE_0:def 3; then A5: a in dom Result s1 & not a in dom (s | A) by RELAT_1:90,SCMFSA_2:66; s +* Initialized I = s0 +* (I +* Start-At insloc 0) by SCMFSA8A:13; hence IExec(I,s).a = (Result s1 +* s | A).a by SCMFSA6B:def 1 .= (Result s1).a by A5,FUNCT_4:12 .= (Computation s1).(LifeSpan s1).a by A4,SCMFSA6B:16 .= s0.a by A1,A3,SCMFSA7B:27 .= (Computation (s0 +* (I +* Start-At insloc 0))).k.a by A1,A3,SCMFSA7B:27; end; theorem Th90: ::TMP29 for s being State of SCM+FSA, I being parahalting Macro-Instruction, a being Int-Location,k being Nat st I does_not_destroy a holds IExec(I,s).a = (Computation (Initialize s +* (I +* Start-At insloc 0))).k.a proof let s be State of SCM+FSA; let I be parahalting Macro-Instruction; let a be Int-Location; let k be Nat; assume A1: I does_not_destroy a; set s0 = Initialize s; set s1 = s0 +* (I +* Start-At insloc 0); A2: I is_closed_on s0 by SCMFSA7B:24; I +* Start-At insloc 0 c= s1 by FUNCT_4:26; then s1 is halting by SCMFSA6B:18; then I is_halting_on s0 by SCMFSA7B:def 8; hence thesis by A1,A2,Th89; end; theorem Th91: ::TMP29'' for s being State of SCM+FSA, I being parahalting Macro-Instruction, a being Int-Location st I does_not_destroy a holds IExec(I,s).a = (Initialize s).a proof let s be State of SCM+FSA; let I be parahalting Macro-Instruction; let a be Int-Location; assume A1: I does_not_destroy a; A2: (Initialize s) | D = (Initialize s +* (I +* Start-At insloc 0)) | D by SCMFSA8A:11; thus IExec(I,s).a = (Computation (Initialize s +* (I +* Start-At insloc 0))).0.a by A1,Th90 .= (Initialize s +* (I +* Start-At insloc 0)).a by AMI_1:def 19 .= (Initialize s).a by A2,SCMFSA6A:38; end; theorem Th92: ::T200724'' for s being State of SCM+FSA, I being keeping_0 Macro-Instruction st I is_halting_on Initialize s holds IExec(I,s).intloc 0 = 1 & for k being Nat holds (Computation (Initialize s +* (I +* Start-At insloc 0))).k.intloc 0 = 1 proof let s be State of SCM+FSA; let I be keeping_0 Macro-Instruction; set s0 = Initialize s; set s1 = s0 +* (I +* Start-At insloc 0); set a = intloc 0; assume I is_halting_on s0; then A1: s1 is halting by SCMFSA7B:def 8; not a in A by SCMFSA_2:84; then not a in dom s /\ A by XBOOLE_0:def 3; then A2: a in dom Result s1 & not a in dom (s | A) by RELAT_1:90,SCMFSA_2:66; A3: s0 | D = s1 | D by SCMFSA8A:11; A4: I +* Start-At insloc 0 c= s1 by FUNCT_4:26; A5: now let k be Nat; thus (Computation s1).k.a = s1.a by A4,SCMFSA6B:def 4 .= s0.a by A3,SCMFSA6A:38 .= 1 by SCMFSA6C:3; end; s +* Initialized I = s1 by SCMFSA8A:13; hence IExec(I,s).a = (Result s1 +* s | A).a by SCMFSA6B:def 1 .= (Result s1).a by A2,FUNCT_4:12 .= (Computation s1).(LifeSpan s1).a by A1,SCMFSA6B:16 .= 1 by A5; let k be Nat; thus (Computation s1).k.a = 1 by A5; end; theorem Th93: ::TQ9 for s being State of SCM+FSA, I being Macro-Instruction, a being Int-Location st I does_not_destroy a holds for k being Nat st IC (Computation (s +* (I +* Start-At insloc 0))).k in dom I holds (Computation (s +* (I +* Start-At insloc 0))).(k + 1).a = (Computation (s +* (I +* Start-At insloc 0))).k.a proof let s be State of SCM+FSA; let I be Macro-Instruction; let a be Int-Location; assume A1: I does_not_destroy a; let k be Nat; assume A2: IC (Computation (s +* (I +* Start-At insloc 0))).k in dom I; set s1 = s +* (I +* Start-At insloc 0); A3: I +* Start-At insloc 0 c= s1 by FUNCT_4:26; dom I misses dom Start-At insloc 0 by SF_MASTR:64; then I c= I +* Start-At insloc 0 by FUNCT_4:33; then A4: I c= s1 by A3,XBOOLE_1:1; set l = IC (Computation s1).k; s1.l = I.l by A2,A4,GRFUNC_1:8; then s1.l in rng I by A2,FUNCT_1:def 5; then A5: s1.l does_not_destroy a by A1,SCMFSA7B:def 4; thus (Computation s1).(k + 1).a = (Following (Computation s1).k).a by AMI_1:def 19 .= Exec(CurInstr (Computation s1).k,(Computation s1).k).a by AMI_1:def 18 .= Exec((Computation s1).k.l,(Computation s1).k).a by AMI_1:def 17 .= Exec(s1.l,(Computation s1).k).a by AMI_1:54 .= (Computation s1).k.a by A5,SCMFSA7B:26; end; theorem Th94: ::TQ9' for s being State of SCM+FSA, I being Macro-Instruction, a being Int-Location st I does_not_destroy a holds for m being Nat st (for n being Nat st n < m holds IC (Computation (s +* (I +* Start-At insloc 0))).n in dom I) holds for n being Nat st n <= m holds (Computation (s +* (I +* Start-At insloc 0))).n.a = s.a proof let s be State of SCM+FSA; let I be Macro-Instruction; let a be Int-Location; assume A1: I does_not_destroy a; let m be Nat; assume A2: for n being Nat st n < m holds IC (Computation (s +* (I +* Start-At insloc 0))).n in dom I; let n be Nat; assume A3: n <= m; set s1 = s +* (I +* Start-At insloc 0); defpred P[Nat] means $1 <= m implies (Computation s1).$1.a = s.a; (Computation s1).0.a = s1.a by AMI_1:def 19 .= s.a by Th28; then A4: P[0]; A5: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume A6: P[k]; assume A7: k + 1 <= m; A8: k + 0 < k + 1 by REAL_1:53; then k < m by A7,AXIOMS:22; then IC (Computation s1).k in dom I by A2; hence (Computation s1).(k + 1).a = s.a by A1,A6,A7,A8,Th93,AXIOMS:22; end; for k being Nat holds P[k] from Ind(A4,A5); hence (Computation (s +* (I +* Start-At insloc 0))).n.a = s.a by A3; end; theorem Th95: ::T210921 for s being State of SCM+FSA, I being good Macro-Instruction for m being Nat st (for n being Nat st n < m holds IC (Computation (s +* (I +* Start-At insloc 0))).n in dom I) holds for n being Nat st n <= m holds (Computation (s +* (I +* Start-At insloc 0))).n.intloc 0 = s.intloc 0 proof let s be State of SCM+FSA; let I be good Macro-Instruction; let m be Nat; assume A1: for n being Nat st n < m holds IC (Computation (s +* (I +* Start-At insloc 0))).n in dom I; let n be Nat; assume A2: n <= m; I does_not_destroy intloc 0 by SCMFSA7B:def 5; hence (Computation (s +* (I +* Start-At insloc 0))).n.intloc 0 = s.intloc 0 by A1,A2,Th94; end; theorem Th96: ::T22842 for s being State of SCM+FSA, I being good Macro-Instruction st I is_halting_on Initialize s & I is_closed_on Initialize s holds IExec(I,s).intloc 0 = 1 & for k being Nat holds (Computation (Initialize s +* (I +* Start-At insloc 0))).k.intloc 0 = 1 proof let s be State of SCM+FSA; let I be good Macro-Instruction; set s0 = Initialize s; set s1 = s0 +* (I +* Start-At insloc 0); set a = intloc 0; assume I is_halting_on s0; then A1: s1 is halting by SCMFSA7B:def 8; assume A2: I is_closed_on s0; defpred P[Nat] means for n being Nat st n <= $1 holds (Computation s1).n.intloc 0 = s0.intloc 0; A3: P[0] proof let n be Nat; assume A4: n <= 0; for i being Nat st i < 0 holds IC (Computation s1).i in dom I by A2, SCMFSA7B:def 7; hence (Computation s1).n.intloc 0 = s0.intloc 0 by A4,Th95; end; A5: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume P[k]; let n be Nat; assume A6: n <= k + 1; for i being Nat st i < k + 1 holds IC (Computation s1).i in dom I by A2 ,SCMFSA7B:def 7; hence (Computation s1).n.intloc 0 = s0.intloc 0 by A6,Th95; end; A7: for k being Nat holds P[k] from Ind(A3,A5); A8: now let k be Nat; thus (Computation s1).k.intloc 0 = s0.intloc 0 by A7 .= 1 by SCMFSA6C:3; end; not a in A by SCMFSA_2:84; then not a in dom s /\ A by XBOOLE_0:def 3; then A9: a in dom Result s1 & not a in dom (s | A) by RELAT_1:90,SCMFSA_2:66; s +* Initialized I = s1 by SCMFSA8A:13; hence IExec(I,s).a = (Result s1 +* s | A).a by SCMFSA6B:def 1 .= (Result s1).a by A9,FUNCT_4:12 .= (Computation s1).(LifeSpan s1).a by A1,SCMFSA6B:16 .= 1 by A8; thus thesis by A8; end; theorem for s being State of SCM+FSA, I being good Macro-Instruction st I is_closed_on s holds for k being Nat holds (Computation (s +* (I +* Start-At insloc 0))).k.intloc 0 = s.intloc 0 proof let s be State of SCM+FSA; let I be good Macro-Instruction; assume A1: I is_closed_on s; let k be Nat; I does_not_destroy intloc 0 by SCMFSA7B:def 5; hence (Computation (s +* (I +* Start-At insloc 0))).k.intloc 0 = s.intloc 0 by A1,SCMFSA7B:27; end; theorem Th98: ::TMP27 for s being State of SCM+FSA, I being keeping_0 parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a holds (Computation (Initialize s +* (I ';' SubFrom(a,intloc 0) +* Start-At insloc 0))).(LifeSpan (Initialize s +* (I ';' SubFrom(a,intloc 0) +* Start-At insloc 0))).a = s.a - 1 proof let s be State of SCM+FSA; let I be keeping_0 parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; set s0 = Initialize s; set s1 = s0 +* (I ';' SubFrom(a,intloc 0) +* Start-At insloc 0); A2: a in dom s0 & not a in dom (I +* Start-At insloc 0) by SCMFSA6B:12,SCMFSA_2:66; IExec(I ';' SubFrom(a,intloc 0),s).a = Exec(SubFrom(a,intloc 0),IExec(I,s)).a by SCMFSA6C:7 .= IExec(I,s).a - IExec(I,s).intloc 0 by SCMFSA_2:91 .= IExec(I,s).a - 1 by SCMFSA6B:35 .= (Computation (s0 +* (I +* Start-At insloc 0))).0.a - 1 by A1,Th90 .= (s0 +* (I +* Start-At insloc 0)).a - 1 by AMI_1:def 19 .= s0.a - 1 by A2,FUNCT_4:12; hence (Computation s1).(LifeSpan s1).a = s0.a - 1 by Th88 .= s.a - 1 by SCMFSA6C:3; end; theorem Th99: ::T211205 for i being Instruction of SCM+FSA st i does_not_destroy intloc 0 holds Macro i is good proof let i be Instruction of SCM+FSA; assume i does_not_destroy intloc 0; then Macro i does_not_destroy intloc 0 by Th77; hence Macro i is good by SCMFSA7B:def 5; end; theorem Th100: ::T13' 6B for s1,s2 being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s1 & I is_halting_on s1 & s1 | D = s2 | D holds for k being Nat holds (Computation (s1 +* (I +* Start-At insloc 0))).k, (Computation (s2 +* (I +* Start-At insloc 0))).k equal_outside the Instruction-Locations of SCM+FSA & CurInstr (Computation (s1 +* (I +* Start-At insloc 0))).k = CurInstr (Computation (s2 +* (I +* Start-At insloc 0))).k proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; assume A1: I is_closed_on s1; assume A2: I is_halting_on s1; assume A3: s1 | D = s2 | D; set ss1 = s1 +* (I +* Start-At insloc 0); set ss2 = s2 +* (I +* Start-At insloc 0); I +* Start-At insloc 0 c= ss1 & I c= I +* Start-At insloc 0 by FUNCT_4:26,SCMFSA8A:9; then A4: I c= ss1 by XBOOLE_1:1; I +* Start-At insloc 0 c= ss2 & I c= I +* Start-At insloc 0 by FUNCT_4:26,SCMFSA8A:9; then A5: I c= ss2 by XBOOLE_1:1; hereby let k be Nat; A6: ss1,ss2 equal_outside the Instruction-Locations of SCM+FSA by A3,SCMFSA8B:7; I is_closed_on s2 by A1,A2,A3,SCMFSA8B:8; then for m being Nat st m < k holds IC (Computation ss2).m in dom I by SCMFSA7B:def 7; hence (Computation ss1).k, (Computation ss2).k equal_outside the Instruction-Locations of SCM+FSA by A4,A5,A6,SCMFSA6B:21; then A7: IC (Computation ss1).k = IC (Computation ss2).k by SCMFSA6A:29; A8: IC (Computation ss1).k in dom I by A1,SCMFSA7B:def 7; I is_closed_on s2 by A1,A2,A3,SCMFSA8B:8; then A9: IC (Computation ss2).k in dom I by SCMFSA7B:def 7; thus CurInstr (Computation ss2).k = (Computation ss2).k.IC (Computation ss2).k by AMI_1:def 17 .= ss2.IC (Computation ss2).k by AMI_1:54 .= I.IC (Computation ss2).k by A5,A9,GRFUNC_1:8 .= ss1.IC (Computation ss1).k by A4,A7,A8,GRFUNC_1:8 .= (Computation ss1).k.IC (Computation ss1).k by AMI_1:54 .= CurInstr (Computation ss1).k by AMI_1:def 17; end; end; theorem Th101: ::T14' 6B for s1,s2 being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s1 & I is_halting_on s1 & s1 | D = s2 | D holds LifeSpan (s1 +* (I +* Start-At insloc 0)) = LifeSpan (s2 +* (I +* Start-At insloc 0)) & Result (s1 +* (I +* Start-At insloc 0)), Result (s2 +* (I +* Start-At insloc 0)) equal_outside the Instruction-Locations of SCM+FSA proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; assume A1: I is_closed_on s1; assume A2: I is_halting_on s1; assume A3: s1 | D = s2 | D; set ss1 = s1 +* (I +* Start-At insloc 0); set ss2 = s2 +* (I +* Start-At insloc 0); A4: ss1 is halting by A2,SCMFSA7B:def 8; A5: now let l be Nat; assume A6: CurInstr (Computation ss2).l = halt SCM+FSA; CurInstr (Computation ss1).l = CurInstr (Computation ss2).l by A1,A2,A3,Th100; hence LifeSpan ss1 <= l by A4,A6,SCM_1:def 2; end; A7: CurInstr (Computation ss2).LifeSpan ss1 = CurInstr (Computation ss1).LifeSpan ss1 by A1,A2,A3,Th100 .= halt SCM+FSA by A4,SCM_1:def 2; I is_halting_on s2 by A1,A2,A3,SCMFSA8B:8; then A8: ss2 is halting by SCMFSA7B:def 8; hence LifeSpan ss1 = LifeSpan ss2 by A5,A7,SCM_1:def 2; then A9: Result ss2 = (Computation ss2).LifeSpan ss1 by A8,SCMFSA6B:16; Result ss1 = (Computation ss1).LifeSpan ss1 by A4,SCMFSA6B:16; hence Result ss1, Result ss2 equal_outside the Instruction-Locations of SCM+FSA by A1,A2,A3,A9,Th100; end; canceled; theorem Th103: ::T3829 for s1,s2 being State of SCM+FSA,I being Macro-Instruction st I is_closed_on s1 & I is_halting_on s1 & I +* Start-At insloc 0 c= s1 & I +* Start-At insloc 0 c= s2 & ex k being Nat st (Computation s1).k,s2 equal_outside the Instruction-Locations of SCM+FSA holds Result s1,Result s2 equal_outside the Instruction-Locations of SCM+FSA proof let s1,s2 be State of SCM+FSA; let I be Macro-Instruction; assume A1: I is_closed_on s1; assume A2: I is_halting_on s1; assume A3: I +* Start-At insloc 0 c= s1; assume A4: I +* Start-At insloc 0 c= s2; given k being Nat such that A5: (Computation s1).k,s2 equal_outside A; set s3 = (Computation s1).k; A6: s2 = s2 +* (I +* Start-At insloc 0) by A4,AMI_5:10; IC s3 = IC s2 by A5,SCMFSA8A:6 .= IC (s2 +* I +* Start-At insloc 0) by A6,FUNCT_4:15 .= insloc 0 by AMI_5:79; then IC SCM+FSA in dom s3 & s3.IC SCM+FSA = insloc 0 by AMI_1:def 15,AMI_5:25; then IC SCM+FSA .--> insloc 0 c= s3 by SCMFSA6A:7; then A7: Start-At insloc 0 c= s3 by AMI_3:def 12; A8: s3 | D = s2 | D by A5,SCMFSA8A:6; I c= I +* Start-At insloc 0 by SCMFSA8A:9; then I c= s1 by A3,XBOOLE_1:1; then I c= s3 by AMI_3:43; then I +* Start-At insloc 0 c= s3 by A7,SCMFSA6A:6; then A9: s3 = s3 +* (I +* Start-At insloc 0) by AMI_5:10; A10: s1 = s1 +* (I +* Start-At insloc 0) by A3,AMI_5:10; now let n be Nat; IC (Computation s3).n = IC (Computation s1).(k + n) by AMI_1:51; hence IC (Computation s3).n in dom I by A1,A10,SCMFSA7B:def 7; end; then A11: I is_closed_on s3 by A9,SCMFSA7B:def 7; A12: s1 is halting by A2,A10,SCMFSA7B:def 8; then consider n being Nat such that A13: CurInstr (Computation s1).n = halt SCM+FSA by AMI_1:def 20; A14: k + n >= n by NAT_1:29; CurInstr (Computation s3).n = CurInstr (Computation s1).(k + n) by AMI_1: 51 .= CurInstr (Computation s1).n by A13,A14,AMI_1:52; then s3 is halting by A13,AMI_1:def 20; then I is_halting_on s3 by A9,SCMFSA7B:def 8; then A15: Result s3,Result s2 equal_outside A by A6,A8,A9,A11,Th101; consider k being Nat such that A16: CurInstr (Computation s1).k = halt SCM+FSA by A12,AMI_1:def 20; s1.IC (Computation s1).k = (Computation s1).k.IC (Computation s1).k by AMI_1:54 .= halt SCM+FSA by A16,AMI_1:def 17; hence Result s1,Result s2 equal_outside A by A15,AMI_1:57; end; begin :: loop definition let I be Macro-Instruction, k be Nat; cluster IncAddr(I,k) -> initial programmed; correctness proof now let m,n be Nat; assume A1: insloc n in dom IncAddr(I,k); assume A2: m < n; dom IncAddr(I,k) = dom I by SCMFSA_4:def 6; hence insloc m in dom IncAddr(I,k) by A1,A2,SCMFSA_4:def 4; end; hence thesis by SCMFSA_4:def 4; end; end; definition let I be Macro-Instruction; canceled 3; func loop I -> halt-free Macro-Instruction equals :Def4: ::D17 ((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0)) * I; coherence proof Directed(I,insloc 0) is halt-free Macro-Instruction; hence thesis by SCMFSA8A:def 1; end; end; theorem Th104: ::T1 for I being Macro-Instruction holds loop I = Directed(I,insloc 0) proof let I be Macro-Instruction; thus loop I = ((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0)) * I by Def4 .= Directed(I,insloc 0) by SCMFSA8A:def 1; end; theorem for I being Macro-Instruction, a being Int-Location holds I does_not_destroy a implies loop I does_not_destroy a proof let I be Macro-Instruction; let a be Int-Location; assume A1: I does_not_destroy a; loop I = Directed(I,insloc 0) by Th104; hence thesis by A1,SCMFSA8A:27; end; definition let I be good Macro-Instruction; cluster loop I -> good; correctness proof loop I = Directed(I,insloc 0) by Th104; hence thesis; end; end; theorem Th106: ::SCMFSA6A'14 for I being Macro-Instruction holds dom loop I = dom I proof let I be Macro-Instruction; thus dom loop I = dom Directed(I,insloc 0) by Th104 .= dom I by SCMFSA8A:19; end; theorem Th107: ::SCMFSA6A'18 for I being Macro-Instruction holds not halt SCM+FSA in rng loop I proof let I be Macro-Instruction; A1: halt SCM+FSA <> goto insloc 0 by SCMFSA_2:47,124; loop I = ((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0))*I by Def4; then rng loop I c= rng((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0)) by RELAT_1:45; hence not halt SCM+FSA in rng loop I by A1,FUNCT_7:14; end; theorem Th108: ::SCMFSA6A'54 for I being Macro-Instruction, x being set holds I.x <> halt SCM+FSA implies (loop I).x = I.x proof let I be Macro-Instruction; let x be set; assume I.x <> halt SCM+FSA; then not I.x in {halt SCM+FSA} by TARSKI:def 1; then not I.x in dom (halt SCM+FSA .--> goto insloc 0) by CQC_LANG:5; then A1: not x in dom ((halt SCM+FSA .--> goto insloc 0) * I) by FUNCT_1:21; A2: rng I c= the Instructions of SCM+FSA by SCMFSA_4:1; thus (loop I).x = (((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0)) * I).x by Def4 .= (((id the Instructions of SCM+FSA) * I) +* ((halt SCM+FSA .--> goto insloc 0) * I)).x by FUNCT_7:11 .= (I +* ((halt SCM+FSA .--> goto insloc 0) * I)).x by A2,RELAT_1:79 .= I.x by A1,FUNCT_4:12; end; theorem Th109: ::TMP24 for s being State of SCM+FSA, I being Macro-Instruction st I is_closed_on s & I is_halting_on s for m being Nat st m <= LifeSpan (s +* (I +* Start-At insloc 0)) holds (Computation (s +* (I +* Start-At insloc 0))).m, (Computation(s +* (loop I +* Start-At insloc 0))).m equal_outside the Instruction-Locations of SCM+FSA proof let s be State of SCM+FSA; let I be Macro-Instruction; set s1 = s +* (I +* Start-At insloc 0); set s2 = s +* (loop I +* Start-At insloc 0); set C1 = Computation s1; set C2 = Computation s2; assume A1: I is_closed_on s; assume A2: I is_halting_on s; defpred X[Nat] means $1 <= LifeSpan s1 implies (Computation s1).$1,(Computation s2).$1 equal_outside A; A3: X[0] proof assume 0 <= LifeSpan s1; s,s +* I equal_outside A by SCMFSA6A:27; then A4: s +* I,s equal_outside A by FUNCT_7:28; s,s +* loop I equal_outside A by SCMFSA6A:27; then s +* I,s +* loop I equal_outside A by A4,FUNCT_7:29; then s +* I +* Start-At insloc 0,s +* loop I +* Start-At insloc 0 equal_outside A by SCMFSA6A:11; then s +* I +* Start-At insloc 0,s2 equal_outside A by FUNCT_4:15; then s1,s2 equal_outside A by FUNCT_4:15; then s1,(Computation s2).0 equal_outside A by AMI_1:def 19; hence (Computation s1).0,(Computation s2).0 equal_outside A by AMI_1:def 19; end; A5: s1 is halting by A2,SCMFSA7B:def 8; A6: for m being Nat st X[m] holds X[m+1] proof let m be Nat; assume A7: m <= LifeSpan s1 implies C1.m,C2.m equal_outside A; assume A8: m + 1 <= LifeSpan s1; then A9: m < LifeSpan s1 by NAT_1:38; A10: C1.(m + 1) = Following C1.m by AMI_1:def 19 .= Exec(CurInstr C1.m,C1.m) by AMI_1:def 18; A11: C2.(m + 1) = Following C2.m by AMI_1:def 19 .= Exec(CurInstr C2.m,C2.m) by AMI_1:def 18; A12: IC C1.m = IC C2.m by A7,A8,NAT_1:38,SCMFSA6A:29; I +* Start-At insloc 0 c= s1 & I c= I +* Start-At insloc 0 by FUNCT_4:26,SCMFSA8A:9; then I c= s1 by XBOOLE_1:1; then A13: I c= C1.m by AMI_3:38; loop I +* Start-At insloc 0 c= s2 & loop I c= loop I +* Start-At insloc 0 by FUNCT_4:26,SCMFSA8A:9; then loop I c= s2 by XBOOLE_1:1; then A14: loop I c= C2.m by AMI_3:38; A15: IC C1.m in dom I by A1,SCMFSA7B:def 7; then A16: IC C1.m in dom loop I by Th106; A17: CurInstr C1.m = C1.m.IC C1.m by AMI_1:def 17 .= I.IC C1.m by A13,A15,GRFUNC_1:8; then I.IC C1.m <> halt SCM+FSA by A5,A9,SCM_1:def 2; then I.IC C1.m = (loop I).IC C1.m by Th108; then CurInstr C1.m = C2.m.IC C1.m by A14,A16,A17,GRFUNC_1:8 .= CurInstr C2.m by A12,AMI_1:def 17; hence C1.(m + 1),C2.(m + 1) equal_outside A by A7,A8,A10,A11,NAT_1:38, SCMFSA6A:32; end; thus for m being Nat holds X[m] from Ind(A3,A6); end; theorem Th110: ::TMP25 for s being State of SCM+FSA, I being Macro-Instruction st I is_closed_on s & I is_halting_on s for m being Nat st m < LifeSpan (s +* (I +* Start-At insloc 0)) holds CurInstr (Computation (s +* (I +* Start-At insloc 0))).m = CurInstr (Computation(s +* (loop I +* Start-At insloc 0))).m proof let s be State of SCM+FSA; let I be Macro-Instruction; set s1 = s +* (I +* Start-At insloc 0); set s2 = s +* (loop I +* Start-At insloc 0); set C1 = Computation s1; set C2 = Computation s2; assume A1: I is_closed_on s & I is_halting_on s; let m be Nat; assume A2: m < LifeSpan (s +* (I +* Start-At insloc 0)); then (Computation (s +* (I +* Start-At insloc 0))).m, (Computation(s +* (loop I +* Start-At insloc 0))).m equal_outside the Instruction-Locations of SCM+FSA by A1,Th109; then A3: IC C1.m = IC C2.m by SCMFSA6A:29; I +* Start-At insloc 0 c= s1 & I c= I +* Start-At insloc 0 by FUNCT_4:26,SCMFSA8A:9; then I c= s1 by XBOOLE_1:1; then A4: I c= C1.m by AMI_3:38; loop I +* Start-At insloc 0 c= s2 & loop I c= loop I +* Start-At insloc 0 by FUNCT_4:26,SCMFSA8A:9; then loop I c= s2 by XBOOLE_1:1; then A5: loop I c= C2.m by AMI_3:38; A6: IC C1.m in dom I by A1,SCMFSA7B:def 7; then A7: IC C1.m in dom loop I by Th106; A8: s1 is halting by A1,SCMFSA7B:def 8; A9: CurInstr C1.m = C1.m.IC C1.m by AMI_1:def 17 .= I.IC C1.m by A4,A6,GRFUNC_1:8; then I.IC C1.m <> halt SCM+FSA by A2,A8,SCM_1:def 2; then I.IC C1.m = (loop I).IC C1.m by Th108; hence CurInstr C1.m = C2.m.IC C1.m by A5,A7,A9,GRFUNC_1:8 .= CurInstr C2.m by A3,AMI_1:def 17; end; Lm6: for s being State of SCM+FSA, I being Macro-Instruction st I is_closed_on s & I is_halting_on s holds (CurInstr (Computation (s +* (loop I +* Start-At insloc 0))). LifeSpan (s +* (I +* Start-At insloc 0)) = goto insloc 0 & for m being Nat st m <= LifeSpan (s +* (I +* Start-At insloc 0)) holds CurInstr (Computation (s +* (loop I +* Start-At insloc 0))).m <> halt SCM+FSA) proof let s be State of SCM+FSA; let I be Macro-Instruction; set s1 = s +* (I +* Start-At insloc 0); set s2 = s +* (loop I +* Start-At insloc 0); set C1 = Computation s1; set C2 = Computation s2; assume A1: I is_closed_on s & I is_halting_on s; then A2: s1 is halting by SCMFSA7B:def 8; set k = LifeSpan s1; A3: CurInstr C1.k = halt SCM+FSA by A2,SCM_1:def 2; C1.k,C2.k equal_outside A by A1,Th109; then A4: IC C1.k = IC C2.k by SCMFSA6A:29; A5: not IC C1.k in dom Start-At insloc 0 by SCMFSA6B:11; A6: IC C1.k in dom I by A1,SCMFSA7B:def 7; then A7: IC C1.k in dom (I +* Start-At insloc 0) by FUNCT_4:13; A8: now thus CurInstr C1.k = C1.k.IC C1.k by AMI_1:def 17 .= s1.IC C1.k by AMI_1:54 .= (I +* Start-At insloc 0).IC C1.k by A7,FUNCT_4:14 .= I.IC C1.k by A5,FUNCT_4:12; end; dom loop I = dom I by Th106; then A9: IC C1.k in dom (loop I +* Start-At insloc 0) by A6,FUNCT_4:13; {halt SCM+FSA} = dom (halt SCM+FSA .--> goto insloc 0) by CQC_LANG:5; then A10: halt SCM+FSA in dom (halt SCM+FSA .--> goto insloc 0) by TARSKI:def 1; A11: (halt SCM+FSA .--> goto insloc 0).halt SCM+FSA = goto insloc 0 by CQC_LANG:6; thus A12: now thus CurInstr C2.LifeSpan s1 = C2.k.IC C1.k by A4,AMI_1:def 17 .= s2.IC C1.k by AMI_1:54 .= (loop I +* Start-At insloc 0).IC C1.k by A9,FUNCT_4:14 .= (loop I).IC C1.k by A5,FUNCT_4:12 .= (((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0))* I ).IC C1.k by Def4 .= ((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0)).halt SCM+FSA by A3,A6,A8,FUNCT_1:23 .= goto insloc 0 by A10,A11,FUNCT_4:14; end; let m be Nat; assume A13: m <= LifeSpan s1; per cases by A13,REAL_1:def 5; suppose A14: m < LifeSpan s1; then CurInstr C1.m <> halt SCM+FSA by A2,SCM_1:def 2; hence CurInstr C2.m <> halt SCM+FSA by A1,A14,Th110; suppose m = LifeSpan s1; hence CurInstr C2.m <> halt SCM+FSA by A12,SCMFSA_2:47,124; end; theorem for s being State of SCM+FSA, I being Macro-Instruction st I is_closed_on s & I is_halting_on s for m being Nat st m <= LifeSpan (s +* (I +* Start-At insloc 0)) holds CurInstr (Computation (s +* (loop I +* Start-At insloc 0))).m <> halt SCM+FSA by Lm6; theorem ::TMP26 for s being State of SCM+FSA, I being Macro-Instruction st I is_closed_on s & I is_halting_on s holds CurInstr (Computation (s +* (loop I +* Start-At insloc 0))). LifeSpan (s +* (I +* Start-At insloc 0)) = goto insloc 0 by Lm6; theorem Th113: ::MAI1 for s being State of SCM+FSA, I being paraclosed Macro-Instruction st I +* Start-At insloc 0 c= s & s is halting for m being Nat st m <= LifeSpan s holds (Computation s).m,(Computation (s +* loop I)).m equal_outside the Instruction-Locations of SCM+FSA proof let s be State of SCM+FSA; let I be paraclosed Macro-Instruction; assume A1: I +* Start-At insloc 0 c= s; assume A2: s is halting; defpred X[Nat] means $1 <= LifeSpan s implies (Computation s).$1,(Computation(s+*loop I)).$1 equal_outside the Instruction-Locations of SCM+FSA; (Computation s).0 = s & (Computation(s+*loop I)).0 = s+*loop I by AMI_1:def 19; then A3: X[0] by SCMFSA6A:27; A4: for m st X[m] holds X[m+1] proof let m; assume A5: m <= LifeSpan s implies (Computation s).m,(Computation(s+*loop I)).m equal_outside the Instruction-Locations of SCM+FSA; assume A6: m+1 <= LifeSpan s; then A7: m < LifeSpan s by NAT_1:38; set Cs = Computation s, CsIJ = Computation(s+*loop I); A8: Cs.(m+1) = Following Cs.m by AMI_1:def 19 .= Exec(CurInstr Cs.m,Cs.m) by AMI_1:def 18; A9: CsIJ.(m+1) = Following CsIJ.m by AMI_1:def 19 .= Exec(CurInstr CsIJ.m,CsIJ.m) by AMI_1:def 18; A10: IC(Cs.m) = IC(CsIJ.m) by A5,A6,NAT_1:38,SCMFSA6A:29; A11: IC Cs.m in dom I by A1,SCMFSA6B:def 2; dom I misses dom Start-At insloc 0 by SF_MASTR:64; then I c= I +* Start-At insloc 0 by FUNCT_4:33; then I c= s by A1,XBOOLE_1:1; then A12: I c= Cs.m by AMI_3:38; loop I c= s+*loop I by FUNCT_4:26; then A13: loop I c= CsIJ.m by AMI_3:38; A14: IC Cs.m in dom loop I by A11,Th106; A15: CurInstr(Cs.m) = (Cs.m).IC(Cs.m) by AMI_1:def 17 .= I.IC(Cs.m) by A11,A12,GRFUNC_1:8; then I.IC(Cs.m) <> halt SCM+FSA by A2,A7,SCM_1:def 2; then CurInstr(Cs.m) = (loop I).IC(Cs.m) by A15,Th108 .= (CsIJ.m).IC(Cs.m) by A13,A14,GRFUNC_1:8 .= CurInstr(CsIJ.m) by A10,AMI_1:def 17; hence (Computation s).(m+1),(Computation(s+*loop I)).(m+1) equal_outside the Instruction-Locations of SCM+FSA by A5,A6,A8,A9,NAT_1:38 ,SCMFSA6A:32; end; thus for m being Nat holds X[m] from Ind(A3,A4); end; theorem for s being State of SCM+FSA, I being parahalting Macro-Instruction st Initialized I c= s holds for k being Nat st k <= LifeSpan s holds CurInstr (Computation (s +* loop I)).k <> halt SCM+FSA proof let s be State of SCM+FSA; let I be parahalting Macro-Instruction; assume A1: Initialized I c= s; then A2: s is halting by AMI_1:def 26; A3: I +* Start-At insloc 0 c= s by A1,SCMFSA6B:8; set s2 = s +* loop I; set m = LifeSpan s; set A = the Instruction-Locations of SCM+FSA; A4: (Computation s).m, (Computation s2).m equal_outside A by A2,A3,Th113 ; set l1 = IC (Computation s).m; A5: IC (Computation s).m in dom I by A3,SCMFSA6B:def 2; then IC (Computation s2).m in dom I by A4,SCMFSA6A:29; then A6: IC (Computation s2).m in dom loop I by Th106; A7: l1 = IC (Computation s2).m by A4,SCMFSA6A:29; I c= Initialized I by SCMFSA6A:26; then dom I c= dom Initialized I by GRFUNC_1:8; then s.l1 = (Initialized I).l1 by A1,A5,GRFUNC_1:8; then A8: I.l1 = s.l1 by A5,SCMFSA6A:50 .= (Computation s).m.IC (Computation s).m by AMI_1:54 .= CurInstr (Computation s).m by AMI_1:def 17 .= halt SCM+FSA by A2,SCM_1:def 2; {halt SCM+FSA} = dom (halt SCM+FSA .--> goto insloc 0) by CQC_LANG:5; then A9: halt SCM+FSA in dom (halt SCM+FSA .--> goto insloc 0) by TARSKI:def 1; A10: (halt SCM+FSA .--> goto insloc 0).halt SCM+FSA = goto insloc 0 by CQC_LANG:6; A11: s2.l1 = (loop I).l1 by A6,A7,FUNCT_4:14 .= (((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0))*I).l1 by Def4 .= ((id the Instructions of SCM+FSA) +* (halt SCM+FSA .--> goto insloc 0)).(halt SCM+FSA) by A5,A8,FUNCT_1:23 .= goto insloc 0 by A9,A10,FUNCT_4:14; A12: CurInstr (Computation s2).m = (Computation s2).m.l1 by A7,AMI_1:def 17 .= goto insloc 0 by A11,AMI_1:54; hereby let k be Nat; assume A13: k <= LifeSpan s; set lk = IC (Computation s).k; per cases; suppose k <> LifeSpan s; A14: (Computation s).k, (Computation s2).k equal_outside A by A2,A3,A13,Th113 ; A15: IC (Computation s).k in dom I by A3,SCMFSA6B:def 2; A16: lk = IC (Computation s2).k by A14,SCMFSA6A:29; A17: dom I = dom loop I by Th106; assume A18: CurInstr (Computation (s +* loop I)).k = halt SCM+FSA; A19: CurInstr (Computation s2).k = (Computation s2).k.lk by A16,AMI_1:def 17 .= s2.lk by AMI_1:54 .= (loop I).lk by A15,A17,FUNCT_4:14; (loop I).lk in rng loop I by A15,A17,FUNCT_1:def 5; hence contradiction by A18,A19,Th107; suppose A20: k = LifeSpan s; assume CurInstr (Computation (s +* loop I)).k = halt SCM+FSA; hence contradiction by A12,A20,SCMFSA_2:47,124; end; end; begin :: Times definition let a be Int-Location; let I be Macro-Instruction; func Times(a,I) -> Macro-Instruction equals :Def5: ::D13 if>0(a,loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)), SCM+FSA-Stop); correctness; end; theorem Th115: ::T211147 *** n.t for I being good Macro-Instruction, a being read-write Int-Location holds if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is good proof let I be good Macro-Instruction; let a be read-write Int-Location; SubFrom(a,intloc 0) does_not_destroy intloc 0 by SCMFSA7B:14; then reconsider J3 = Macro SubFrom(a,intloc 0) as good Macro-Instruction by Th99; I ';' SubFrom(a,intloc 0) = I ';' J3 by SCMFSA6A:def 6; then reconsider I1 = I ';' SubFrom(a,intloc 0) as good Macro-Instruction; a =0_goto insloc (card I1 + 3) does_not_destroy intloc 0 by SCMFSA7B:18; then reconsider J1 = Macro (a =0_goto insloc (card I1 + 3)) as good Macro-Instruction by Th99; if=0(a,Goto insloc 2,I1) = (a =0_goto insloc (card I1 + 3) ';' I1 ';' Goto insloc (card Goto insloc 2 + 1)) ';' Goto insloc 2 ';' SCM+FSA-Stop by SCMFSA8B:def 1 .= J1 ';' I1 ';' Goto insloc (card Goto insloc 2 + 1) ';' Goto insloc 2 ';' SCM+FSA-Stop by SCMFSA6A:def 5; hence if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is good; end; theorem Th116: ::T21921 ** n.t for I,J being Macro-Instruction,a being Int-Location holds if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)). insloc (card (I ';' SubFrom(a,intloc 0)) + 3) = goto insloc (card (I ';' SubFrom(a,intloc 0)) + 5) proof let I,J be Macro-Instruction; let a be Int-Location; set I1 = I ';' SubFrom(a,intloc 0); set J3 = a =0_goto insloc (card I1 + 3) ';' I1 ';' Goto insloc (card Goto insloc 2 + 1); set J4 = (a =0_goto insloc (card I1 + 3) ';' I1 ';' Goto insloc (card Goto insloc 2 + 1)) ';' Goto insloc 2; A1: card (a =0_goto insloc (card I1 + 3) ';' I1) = card (Macro (a =0_goto insloc (card I1 + 3)) ';' I1) by SCMFSA6A:def 5 .= card Macro (a =0_goto insloc (card I1 + 3)) + card I1 by SCMFSA6A:61 .= 2 + card I1 by SCMFSA7B:6; card Goto insloc (card Goto insloc 2 + 1) = 1 by SCMFSA8A:29; then A2: card J3 = card I1 + 2 + 1 by A1,SCMFSA6A:61 .= card I1 + (2 + 1) by XCMPLX_1:1; A3: card Goto insloc 2 = 1 by SCMFSA8A:29; then A4: card J4 = card I1 + (2 + 1) + 1 by A2,SCMFSA6A:61 .= card I1 + (2 + 1 + 1) by XCMPLX_1:1; card I1 + (2 + 1) -' card J3 = 0 by A2,GOBOARD9:1; then A5: goto insloc 2 = (Goto insloc 2).insloc (card I1 + (2 + 1) -' card J3) by SCMFSA8A:47; card I1 + (2 + 1) < card J3 + card Goto insloc 2 by A2,A3,NAT_1:38; then A6: J4.insloc (card I1 + (2 + 1)) = IncAddr(goto insloc 2,card J3) by A2,A5,Th13 .= goto (insloc 2 + (card I1 + (2 + 1))) by A2,SCMFSA_4:14 .= goto insloc (2 + (card I1 + (2 + 1))) by SCMFSA_4:def 1 .= goto insloc (card I1 + (2 + 3)) by XCMPLX_1:1; A7: 6 <> 0 & InsCode goto insloc (card I1 + 5) = 6 & InsCode halt SCM+FSA = 0 by SCMFSA_2:47,124; card J4 = card I1 + 3 + 1 by A4,XCMPLX_1:1; then card I1 + 3 < card J4 by NAT_1:38; then A8: insloc (card I1 + 3) in dom J4 by SCMFSA6A:15; then (J4 ';' SCM+FSA-Stop).insloc (card I1 + 3) = (Directed J4).insloc (card I1 + 3) by SCMFSA8A:28 .= goto insloc (card I1 + 5) by A6,A7,A8,SCMFSA8A:30; hence if=0(a,Goto insloc 2,I1).insloc (card I1 + 3) = goto insloc (card I1 + 5) by SCMFSA8B:def 1; end; theorem Th117: ::TMP22 for s being State of SCM+FSA, I being good parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a & s.intloc 0 = 1 & s.a > 0 holds loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on s proof let s be State of SCM+FSA; let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; assume A2: s.intloc 0 = 1; assume A3: s.a > 0; set P = if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)); reconsider I1 = I ';' SubFrom(a,intloc 0) as parahalting Macro-Instruction; set i = a =0_goto insloc (card I1 + 3); defpred P[Nat] means for s being State of SCM+FSA st s.intloc 0 = 1 & s.a = $1 & s.a > 0 holds ((Computation (s +* (loop P +* Start-At insloc 0))). (LifeSpan (s +* (P +* Start-At insloc 0)) + 1)).a = s.a - 1 & ((Computation (s +* (loop P +* Start-At insloc 0))). (LifeSpan (s +* (P +* Start-At insloc 0)) + 1)).intloc 0 = 1 & ex k being Nat st IC (Computation (s +* (loop P +* Start-At insloc 0))).k = insloc card ProgramPart loop P & for n being Nat st n < k holds IC (Computation (s +* (loop P +* Start-At insloc 0))).n in dom loop P; A4: P[0]; A5: for k being Nat holds P[k] implies P[k + 1] proof let k be Nat; assume A6: P[k]; let ss be State of SCM+FSA; assume A7: ss.intloc 0 = 1; assume A8: ss.a = k + 1; assume A9: ss.a > 0; set s1 = ss +* (P +* Start-At insloc 0); set s2 = ss +* (loop P +* Start-At insloc 0); set C1 = Computation s1; set C2 = Computation s2; set s3 = C2.(LifeSpan s1 + 1); A10: now A11: now thus card loop P = card dom loop P by PRE_CIRC:21 .= card dom P by Th106 .= card P by PRE_CIRC:21; thus card P = card Goto insloc 2 + card I1 + 4 by SCMFSA8B:14 .= card I1 + 1 + 4 by SCMFSA8A:29 .= card I1 + (4 + 1) by XCMPLX_1:1 .= card I1 + (3 + 2) .= card I1 + 3 + 2 by XCMPLX_1:1; end; A12: now thus P = i ';' I1 ';' Goto insloc (card Goto insloc 2 + 1) ';' Goto insloc 2 ';' SCM+FSA-Stop by SCMFSA8B:def 1 .= i ';' (I1 ';' Goto insloc (card Goto insloc 2 + 1)) ';' Goto insloc 2 ';' SCM+FSA-Stop by SCMFSA6A:66 .= i ';' (I1 ';' Goto insloc (card Goto insloc 2 + 1) ';' Goto insloc 2) ';' SCM+FSA-Stop by SCMFSA6A:66 .= i ';' (I1 ';' Goto insloc (card Goto insloc 2 + 1) ';' Goto insloc 2 ';' SCM+FSA-Stop) by SCMFSA6A:66 .= Macro i ';' (I1 ';' Goto insloc (card Goto insloc 2 + 1) ';' Goto insloc 2 ';' SCM+FSA-Stop) by SCMFSA6A:def 5; end; InsCode i = 7 & InsCode halt SCM+FSA = 0 by SCMFSA_2:48,124 ; then insloc 0 in dom Macro i & (Macro i).insloc 0 <> halt SCM+FSA by SCMFSA6B:32,33; hence P.insloc 0 = (Macro i).insloc 0 by A12,SCMFSA6A:54 .= i by SCMFSA6B:33; hence P.insloc 0 <> halt SCM+FSA by SCMFSA_2:48,124; hereby 0 < 2 & 2 <= card P by A11,NAT_1:29; then 0 < card P; hence insloc 0 in dom P by SCMFSA6A:15; end; card ProgramPart loop P = card loop P by AMI_5:72 .= card I1 + (3 + 2) by A11,XCMPLX_1:1; hence P.insloc (card I1 + 3) = goto insloc card ProgramPart loop P by Th116; hence P.insloc (card I1 + 3) <> halt SCM+FSA by SCMFSA_2:47,124; hereby card I1 + 3 + 0 < card P by A11,REAL_1:53; hence insloc (card I1 + 3) in dom P by SCMFSA6A:15; end; end; A13: now I1 is_closed_on ss & I1 is_halting_on ss by SCMFSA7B:24,25; hence A14: P is_closed_on ss & P is_halting_on ss by A9,SCMFSA8B:18; C2.(LifeSpan s1 + 1) = Following C2.LifeSpan s1 by AMI_1:def 19 .= Exec(CurInstr C2.LifeSpan s1,C2.LifeSpan s1) by AMI_1:def 18; hence A15: C2.(LifeSpan s1 + 1) = Exec(goto insloc 0,C2.LifeSpan s1) by A14,Lm6 ; hereby thus IC C2.(LifeSpan s1 + 1) = Exec(goto insloc 0,C2.LifeSpan s1).IC SCM+FSA by A15,AMI_1:def 15 .= insloc 0 by SCMFSA_2:95; end; end; A16: now A17: now thus card loop P = card dom loop P by PRE_CIRC:21 .= card dom P by Th106 .= card P by PRE_CIRC:21; thus card P = card Goto insloc 2 + card I1 + 4 by SCMFSA8B:14 .= card I1 + 1 + 4 by SCMFSA8A:29 .= card I1 + (4 + 1) by XCMPLX_1:1 .= card I1 + (3 + 2) .= card I1 + 3 + 2 by XCMPLX_1:1; end; hereby 0 < 2 & 2 <= card P by A17,NAT_1:29; then 0 < card P; hence insloc 0 in dom loop P by A17,SCMFSA6A:15; card I1 + 3 + (1 + 1) = (card I1 + 3 + 1) + 1 by XCMPLX_1:1; then card I1 + 3 + 1 < card P by A17,NAT_1:38; then card I1 + 3 < card loop P by A17,NAT_1:38; hence insloc (card I1 + 3) in dom loop P by SCMFSA6A:15; end; thus A18: now intloc 0 in dom ss by SCMFSA_2:66; then A19: ss +* (intloc 0 .--> 1) = ss by A7,Th6; A20: dom P misses dom Start-At insloc 0 by SF_MASTR:64; Initialize ss = ss +* (intloc 0 .--> 1) +* Start-At insloc 0 by SCMFSA6C:def 3; hence Initialize ss +* (P +* Start-At insloc 0) = ss +* Start-At insloc 0 +* (Start-At insloc 0 +* P) by A19,A20,FUNCT_4:36 .= ss +* Start-At insloc 0 +* Start-At insloc 0 +* P by FUNCT_4:15 .= ss +* (Start-At insloc 0 +* Start-At insloc 0) +* P by FUNCT_4: 15 .= ss +* (Start-At insloc 0 +* P) by FUNCT_4:15 .= ss +* (P +* Start-At insloc 0) by A20,FUNCT_4:36; end; consider Is being State of SCM+FSA such that A21: Is = Initialize ss +* (P +* Start-At insloc 0); A22: I1 is_halting_on Initialize ss by SCMFSA7B:25; A23: now A24: now let b be Int-Location; C1.LifeSpan s1,C2.LifeSpan s1 equal_outside A & C2.(LifeSpan s1 + 1).b = C2.(LifeSpan s1).b by A13,Th109,SCMFSA_2:95; hence C2.(LifeSpan s1 + 1).b = (Computation Is).(LifeSpan Is).b by A18,A21,SCMFSA6A:30; end; (Initialize ss).a > 0 & I1 is_closed_on Initialize ss & I1 is_halting_on Initialize ss by A9,SCMFSA6C:3,SCMFSA7B:24,25; then A25: P is_halting_on Initialize ss & P is_closed_on Initialize ss by SCMFSA8B:18; thus C2.(LifeSpan s1 + 1).a = (Computation Is).(LifeSpan Is).a by A24 .= IExec(P,ss).a by A21,A25,Th87; A26: P is good by Th115; thus C2.(LifeSpan s1 + 1).intloc 0 = (Computation Is).(LifeSpan Is).intloc 0 by A24 .= 1 by A21,A25,A26,Th96; end; ss.a <> 0 & I1 is_closed_on Initialize ss by A9,SCMFSA7B:24; then IExec(P,ss) = IExec(I1,ss) +* Start-At insloc (card Goto insloc 2 + card I1 + 3) by A22,SCMFSA8B:19; then IExec(P,ss).a = IExec(I1,ss).a by SCMFSA_3:11; hence C2.(LifeSpan s1 + 1).a = (Computation (Initialize ss +* (I1 +* Start-At insloc 0))). (LifeSpan (Initialize ss +* (I1 +* Start-At insloc 0))).a by A22,A23,Th87 .= ss.a - 1 by A1,Th98; thus C2.(LifeSpan s1 + 1).intloc 0 = 1 by A23; end; hence s3.a = ss.a - 1 & s3.intloc 0 = 1; A27: s3.a = k by A8,A16,XCMPLX_1:26; hereby per cases by NAT_1:19; suppose A28: k = 0; take m = LifeSpan s1 + 1 + 1 + 1; A29: now thus CurInstr C2.(LifeSpan s1 + 1) = C2.(LifeSpan s1 + 1).insloc 0 by A13,AMI_1:def 17 .= s2.insloc 0 by AMI_1:54 .= (loop P).insloc 0 by A16,Th26 .= i by A10,Th108; end; A30: now thus C2.(LifeSpan s1 + 1 + 1) = Following C2.(LifeSpan s1 + 1) by AMI_1:def 19 .= Exec(i,C2.(LifeSpan s1 + 1)) by A29,AMI_1:def 18; end; A31: now thus IC C2.(LifeSpan s1 + 1 + 1) = C2.(LifeSpan s1 + 1 + 1).IC SCM+FSA by AMI_1:def 15 .= insloc (card I1 + 3) by A8,A16,A28,A30,SCMFSA_2:96; end; A32: now thus CurInstr C2.(LifeSpan s1 + 1 + 1) = C2.(LifeSpan s1 + 1 + 1).insloc (card I1 + 3) by A31,AMI_1:def 17 .= s2.insloc (card I1 + 3) by AMI_1:54 .= (loop P).insloc (card I1 + 3) by A16,Th26 .= goto insloc card ProgramPart loop P by A10,Th108; end; A33: C2.m = Following C2.(LifeSpan s1 + 1 + 1) by AMI_1:def 19 .= Exec(goto insloc card ProgramPart loop P,C2.(LifeSpan s1 + 1 + 1)) by A32,AMI_1:def 18; thus IC C2.m = C2.m.IC SCM+FSA by AMI_1:def 15 .= insloc card ProgramPart loop P by A33,SCMFSA_2:95; hereby let n be Nat; assume n < m; then n <= LifeSpan s1 + 1 + 1 by NAT_1:38; then A34: n <= LifeSpan s1 + 1 or n = LifeSpan s1 + 1 + 1 by NAT_1:26; per cases by A34,NAT_1:26; suppose A35: n <= LifeSpan s1; I1 is_closed_on ss & I1 is_halting_on ss by SCMFSA7B:24,25 ; then A36: P is_closed_on ss & P is_halting_on ss by A9,SCMFSA8B:18; then C1.n,C2.n equal_outside A by A35,Th109; then A37: IC C2.n = IC C1.n by SCMFSA8A:6; IC C1.n in dom P by A36,SCMFSA7B:def 7; hence IC C2.n in dom loop P by A37,Th106; suppose n = LifeSpan s1 + 1; hence IC C2.n in dom loop P by A13,A16; suppose n = LifeSpan s1 + 1 + 1; hence IC C2.n in dom loop P by A16,A31; end; suppose A38: k > 0; consider Is3 being State of SCM+FSA such that A39: Is3 = Initialize s3; A40: Is3.intloc 0 = 1 by A39,SCMFSA6C:3; Is3.a = k & Is3.a > 0 by A27,A38,A39,SCMFSA6C:3; then consider m0 being Nat such that A41: IC (Computation (Is3 +* (loop P +* Start-At insloc 0))).m0 = insloc card ProgramPart loop P and A42: for n being Nat st n < m0 holds IC (Computation (Is3 +* (loop P +* Start-At insloc 0))).n in dom loop P by A6,A40; take m = LifeSpan s1 + 1 + m0; A43: now thus A44: s3.IC SCM+FSA = IC s3 by AMI_1:def 15; A45: IC SCM+FSA in dom s3 by AMI_5:25; A46: dom loop P misses dom Start-At insloc 0 by SF_MASTR:64; now thus s2 = ss +* (Start-At insloc 0 +* loop P) by A46,FUNCT_4:36 .= ss +* Start-At insloc 0 +* loop P by FUNCT_4:15; end; then loop P c= s2 by FUNCT_4:26; then ProgramPart loop P c= s3 by AMI_5:64; then A47: loop P c= s3 by AMI_5:72; A48: now thus Initialize s3 +* (loop P +* Start-At insloc 0) = Initialize s3 +* (Start-At insloc 0 +* loop P) by A46,FUNCT_4:36 .= Initialize s3 +* Start-At insloc 0 +* loop P by FUNCT_4:15; end; now thus Initialize s3 +* (loop P +* Start-At insloc 0) = s3 +* Start-At insloc 0 +* loop P by A13,A16,A48,Th14 .= s3 +* (IC SCM+FSA .--> insloc 0) +*loop P by AMI_3:def 12; end; hence Initialize s3 +* (loop P +* Start-At insloc 0) = s3 +* loop P by A13,A44,A45,Th6 .= s3 by A47,AMI_5:10; end; hence IC C2.m = insloc card ProgramPart loop P by A39,A41,AMI_1:51; hereby let n be Nat; assume A49: n < m; I1 is_closed_on ss & I1 is_halting_on ss by SCMFSA7B:24,25 ; then A50: P is_closed_on ss & P is_halting_on ss by A9,SCMFSA8B:18; per cases by NAT_1:38; suppose n <= LifeSpan s1; then C1.n,C2.n equal_outside A by A50,Th109; then A51: IC C2.n = IC C1.n by SCMFSA8A:6; IC C1.n in dom P by A50,SCMFSA7B:def 7; hence IC C2.n in dom loop P by A51,Th106; suppose A52: LifeSpan s1 + 1 <= n; consider mm being Nat such that A53: mm = n -' (LifeSpan s1 + 1); mm + (LifeSpan s1 + 1) = n by A52,A53,AMI_5:4; then A54: IC C2.n = IC (Computation s3).mm by AMI_1:51; n - (LifeSpan s1 + 1) >= 0 by A52,SQUARE_1:12; then mm = n - (LifeSpan s1 + 1) & m0 = m - (LifeSpan s1 + 1) by A53,BINARITH:def 3,XCMPLX_1:26; then mm < m0 by A49,REAL_1:54; hence IC C2.n in dom loop P by A39,A42,A43,A54; end; end; end; reconsider sa = s.a as Nat by A3,INT_1:16; for k being Nat holds P[k] from Ind(A4,A5); then P[sa]; then ex k being Nat st IC (Computation (s +* (loop P +* Start-At insloc 0))).k = insloc card ProgramPart loop P & for n being Nat st n < k holds IC (Computation (s +* (loop P +* Start-At insloc 0))).n in dom loop P by A2,A3; hence loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on s by SCMFSA8A:def 3; end; theorem for s being State of SCM+FSA, I being good parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a & s.a > 0 holds Initialized loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on s proof let s be State of SCM+FSA; let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; assume A2: s.a > 0; (Initialize s).a = s.a & (Initialize s).intloc 0 = 1 by SCMFSA6C:3; then loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on Initialize s by A1,A2,Th117; hence Initialized loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on s by Lm1; end; theorem for s being State of SCM+FSA, I being good parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a & s.intloc 0 = 1 holds Times(a,I) is_closed_on s & Times(a,I) is_halting_on s proof let s be State of SCM+FSA; let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; assume A2: s.intloc 0 = 1; A3: Times(a,I) = if>0(a,loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)), SCM+FSA-Stop) by Def5; per cases; suppose A4: s.a > 0; Directed loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) = loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) by SCMFSA8A:40; then Directed loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on s by A1,A2,A4,Th117; hence Times(a,I) is_closed_on s & Times(a,I) is_halting_on s by A3,A4,Th68; suppose A5: s.a <= 0; SCM+FSA-Stop is_closed_on s & SCM+FSA-Stop is_halting_on s by SCMFSA7B:24,25; hence Times(a,I) is_closed_on s & Times(a,I) is_halting_on s by A3,A5,SCMFSA8B:24; end; theorem for I being good parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a holds Initialized Times(a,I) is halting proof let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; A2: Times(a,I) = if>0(a,loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)), SCM+FSA-Stop) by Def5; now let s be State of SCM+FSA; per cases; suppose s.a > 0; then A3: (Initialize s).intloc 0 = 1 & (Initialize s).a > 0 by SCMFSA6C:3; Directed loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) = loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) by SCMFSA8A:40; then Directed loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) is_pseudo-closed_on Initialize s by A1,A3,Th117; then Times(a,I) is_halting_on Initialize s by A2,A3,Th68; hence Initialized Times(a,I) is_halting_on s by Th22; suppose s.a <= 0; then A4: (Initialize s).a <= 0 by SCMFSA6C:3; SCM+FSA-Stop is_closed_on Initialize s & SCM+FSA-Stop is_halting_on Initialize s by SCMFSA7B:24,25 ; then Times(a,I) is_halting_on Initialize s by A2,A4,SCMFSA8B:24; hence Initialized Times(a,I) is_halting_on s by Th22; end; hence Initialized Times(a,I) is halting by Th24; end; theorem for I,J being Macro-Instruction, a,c being Int-Location st I does_not_destroy c & J does_not_destroy c holds if=0(a,I,J) does_not_destroy c & if>0(a,I,J) does_not_destroy c proof let I,J be Macro-Instruction; let a,c be Int-Location; assume A1: I does_not_destroy c; assume A2: J does_not_destroy c; a =0_goto insloc (card J + 3) does_not_destroy c by SCMFSA7B:18; then A3: a =0_goto insloc (card J + 3) ';' J does_not_destroy c by A2,Th82; Goto insloc (card I + 1) does_not_destroy c by Th86; then a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) does_not_destroy c by A3,Th81; then A4: a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I does_not_destroy c by A1,Th81; A5: SCM+FSA-Stop does_not_destroy c by Th85; if=0(a,I,J) = a =0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 1; hence if=0(a,I,J) does_not_destroy c by A4,A5,Th81; a >0_goto insloc (card J + 3) does_not_destroy c by SCMFSA7B:19; then A6: a >0_goto insloc (card J + 3) ';' J does_not_destroy c by A2,Th82; Goto insloc (card I + 1) does_not_destroy c by Th86; then a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) does_not_destroy c by A6,Th81; then A7: a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I does_not_destroy c by A1,Th81; A8: SCM+FSA-Stop does_not_destroy c by Th85; if>0(a,I,J) = a >0_goto insloc (card J + 3) ';' J ';' Goto insloc (card I + 1) ';' I ';' SCM+FSA-Stop by SCMFSA8B:def 2; hence if>0(a,I,J) does_not_destroy c by A7,A8,Th81; end; theorem Th122: ::TMP22' for s being State of SCM+FSA, I being good parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a & s.intloc 0 = 1 & s.a > 0 holds ex s2 being State of SCM+FSA, k being Nat st s2 = s +* (loop if=0(a,Goto insloc 2, I ';' SubFrom(a,intloc 0)) +* Start-At insloc 0) & k = LifeSpan (s +* (if=0(a,Goto insloc 2, I ';' SubFrom(a,intloc 0)) +* Start-At insloc 0)) + 1 & (Computation s2).k.a = s.a - 1 & (Computation s2).k.intloc 0 = 1 & (for b being read-write Int-Location st b <> a holds (Computation s2).k.b = IExec(I,s).b) & (for f being FinSeq-Location holds (Computation s2).k.f = IExec(I,s).f) & IC (Computation s2).k = insloc 0 & for n being Nat st n <= k holds IC (Computation s2).n in dom loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)) proof let s be State of SCM+FSA; let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; assume A2: s.intloc 0 = 1; assume A3: s.a > 0; set P = if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)); reconsider I1 = I ';' SubFrom(a,intloc 0) as parahalting Macro-Instruction; set s1 = s +* (P +* Start-At insloc 0); take s2 = s +* (loop P +* Start-At insloc 0); take k = LifeSpan s1 + 1; set C1 = Computation s1; set C2 = Computation s2; thus s2 = s +* (loop if=0(a,Goto insloc 2, I ';' SubFrom(a,intloc 0)) +* Start-At insloc 0) & k = LifeSpan (s +* (if=0(a,Goto insloc 2, I ';' SubFrom(a,intloc 0)) +* Start-At insloc 0)) + 1; I1 is_closed_on s & I1 is_halting_on s by SCMFSA7B:24,25; then A4: P is_closed_on s & P is_halting_on s by A3,SCMFSA8B:18; C2.(LifeSpan s1 + 1) = Following C2.LifeSpan s1 by AMI_1:def 19 .= Exec(CurInstr C2.LifeSpan s1,C2.LifeSpan s1) by AMI_1:def 18; then A5: C2.(LifeSpan s1 + 1) = Exec(goto insloc 0,C2.LifeSpan s1) by A4,Lm6; A6: now thus IC C2.(LifeSpan s1 + 1) = Exec(goto insloc 0,C2.LifeSpan s1).IC SCM+FSA by A5,AMI_1:def 15 .= insloc 0 by SCMFSA_2:95; end; A7: now intloc 0 in dom s by SCMFSA_2:66; then A8: s +* (intloc 0 .--> 1) = s by A2,Th6; A9: dom P misses dom Start-At insloc 0 by SF_MASTR:64; Initialize s = s +* (intloc 0 .--> 1) +* Start-At insloc 0 by SCMFSA6C:def 3; hence Initialize s +* (P +* Start-At insloc 0) = s +* Start-At insloc 0 +* (Start-At insloc 0 +* P) by A8,A9,FUNCT_4:36 .= s +* Start-At insloc 0 +* Start-At insloc 0 +* P by FUNCT_4:15 .= s +* (Start-At insloc 0 +* Start-At insloc 0) +* P by FUNCT_4:15 .= s +* (Start-At insloc 0 +* P) by FUNCT_4:15 .= s +* (P +* Start-At insloc 0) by A9,FUNCT_4:36; end; set Is = Initialize s +* (P +* Start-At insloc 0); A10: now let b be Int-Location; C1.LifeSpan s1,C2.LifeSpan s1 equal_outside A & C2.(LifeSpan s1 + 1).b = C2.(LifeSpan s1).b by A4,A5,Th109,SCMFSA_2:95; hence C2.(LifeSpan s1 + 1).b = (Computation Is).(LifeSpan Is).b by A7, SCMFSA6A:30; end; A11: I1 is_halting_on Initialize s by SCMFSA7B:25; (Initialize s).a > 0 & I1 is_closed_on Initialize s & I1 is_halting_on Initialize s by A3,SCMFSA6C:3,SCMFSA7B:24,25; then A12: P is_halting_on Initialize s & P is_closed_on Initialize s by SCMFSA8B:18; A13: now thus C2.(LifeSpan s1 + 1).a = (Computation Is).(LifeSpan Is).a by A10 .= IExec(P,s).a by A12,Th87; A14: P is good by Th115; thus C2.(LifeSpan s1 + 1).intloc 0 = (Computation Is).(LifeSpan Is).intloc 0 by A10 .= 1 by A12,A14,Th96; end; s.a <> 0 & I1 is_closed_on Initialize s by A3,SCMFSA7B:24; then A15: IExec(P,s) = IExec(I1,s) +* Start-At insloc (card Goto insloc 2 + card I1 + 3) by A11,SCMFSA8B:19; then IExec(P,s).a = IExec(I1,s).a by SCMFSA_3:11; hence (Computation s2).k.a = (Computation (Initialize s +* (I1 +* Start-At insloc 0))). (LifeSpan (Initialize s +* (I1 +* Start-At insloc 0))).a by A11,A13,Th87 .= s.a - 1 by A1,Th98; thus (Computation s2).k.intloc 0 = 1 by A13; hereby let b be read-write Int-Location; assume A16: b <> a; thus (Computation s2).k.b = (Computation Is).(LifeSpan Is).b by A10 .= IExec(P,s).b by A12,Th87 .= IExec(I1,s).b by A15,SCMFSA_3:11 .= Exec(SubFrom(a,intloc 0),IExec(I,s)).b by SCMFSA6C:7 .= IExec(I,s).b by A16,SCMFSA_2:91; end; hereby let f be FinSeq-Location; C1.LifeSpan s1,C2.LifeSpan s1 equal_outside A & C2.(LifeSpan s1 + 1).f = C2.(LifeSpan s1).f by A4,A5,Th109,SCMFSA_2:95; hence (Computation s2).k.f = (Computation Is).(LifeSpan Is).f by A7, SCMFSA6A:31 .= IExec(P,s).f by A12,Th87 .= IExec(I1,s).f by A15,SCMFSA_3:12 .= Exec(SubFrom(a,intloc 0),IExec(I,s)).f by SCMFSA6C:8 .= IExec(I,s).f by SCMFSA_2:91; end; thus IC (Computation s2).k = insloc 0 by A6; hereby let n be Nat; assume A17: n <= k; per cases by A17,NAT_1:26; suppose A18: n <= LifeSpan s1; I1 is_closed_on s & I1 is_halting_on s by SCMFSA7B:24,25; then A19: P is_closed_on s & P is_halting_on s by A3,SCMFSA8B:18; then C1.n,C2.n equal_outside A by A18,Th109; then A20: IC C2.n = IC C1.n by SCMFSA8A:6; IC C1.n in dom P by A19,SCMFSA7B:def 7; hence IC (Computation s2).n in dom loop P by A20,Th106; suppose A21: n = LifeSpan s1 + 1; A22: card loop P = card dom loop P by PRE_CIRC:21 .= card dom P by Th106 .= card P by PRE_CIRC:21; card P = card Goto insloc 2 + card I1 + 4 by SCMFSA8B:14 .= card I1 + 1 + 4 by SCMFSA8A:29 .= card I1 + (4 + 1) by XCMPLX_1:1 .= card I1 + (3 + 2) .= card I1 + 3 + 2 by XCMPLX_1:1; then 0 < 2 & 2 <= card P by NAT_1:29; then 0 < card loop P by A22; hence IC (Computation s2).n in dom loop P by A6,A21,SCMFSA6A:15; end; end; theorem Th123: ::T1 for s being State of SCM+FSA, I being good parahalting Macro-Instruction, a being read-write Int-Location st s.intloc 0 = 1 & s.a <= 0 holds IExec(Times(a,I),s) | (Int-Locations \/ FinSeq-Locations) = s | (Int-Locations \/ FinSeq-Locations) proof let s be State of SCM+FSA; let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: s.intloc 0 = 1; assume s.a <= 0; then A2: (Initialize s).a <= 0 by SCMFSA6C:3; set P = if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)); set s0 = Initialize s; A3: Times(a,I) = if>0(a,loop P,SCM+FSA-Stop) by Def5; A4: SCM+FSA-Stop is_closed_on Initialize s & SCM+FSA-Stop is_halting_on Initialize s by SCMFSA7B:24,25; then A5: Times(a,I) is_closed_on Initialize s & Times(a,I) is_halting_on Initialize s by A2,A3,SCMFSA8B:24; SCM+FSA-Stop is_closed_on s0 & SCM+FSA-Stop is_halting_on s0 by SCMFSA7B:24,25; then A6: Directed SCM+FSA-Stop is_pseudo-closed_on s0 by SCMFSA8A:37; A7: s0.intloc 0 = 1 & (for a being read-write Int-Location holds s0.a = s.a) & for f being FinSeq-Location holds s0.f = s.f by SCMFSA6C:3; then A8: IExec(Times(a,I),s) | D = IExec(Times(a,I),s0) | D by A5,Th45 .= IExec(SCM+FSA-Stop ';' SCM+FSA-Stop,s0) | D by A2,A3,A6,A7,Th73; A9: IExec(SCM+FSA-Stop,s0) | D = (Initialize s0 +* Start-At insloc 0) | D by Th38 .= (Initialize s +* Start-At insloc 0) | D by Th15 .= s0 | D by SCMFSA8A:10; IExec(SCM+FSA-Stop ';' SCM+FSA-Stop,s0) = IExec(SCM+FSA-Stop,IExec(SCM+FSA-Stop,s0)) +* Start-At (IC IExec(SCM+FSA-Stop,IExec(SCM+FSA-Stop,s0)) + card SCM+FSA-Stop) by SCMFSA6B:44; hence IExec(Times(a,I),s) | D = IExec(SCM+FSA-Stop,IExec(SCM+FSA-Stop,s0)) | D by A8,SCMFSA8A:10 .= IExec(SCM+FSA-Stop,s0) | D by A4,A7,A9,Th46 .= s | D by A1,A9,Th27; end; theorem Th124: ::T2 for s being State of SCM+FSA, I being good parahalting Macro-Instruction, a being read-write Int-Location st I does_not_destroy a & s.a > 0 holds IExec(I ';' SubFrom(a,intloc 0),s).a = s.a - 1 & IExec(Times(a,I),s) | (Int-Locations \/ FinSeq-Locations) = IExec(Times(a,I),IExec(I ';' SubFrom(a,intloc 0),s)) | (Int-Locations \/ FinSeq-Locations) proof let s be State of SCM+FSA; let I be good parahalting Macro-Instruction; let a be read-write Int-Location; assume A1: I does_not_destroy a; assume A2: s.a > 0; set I1 = I ';' SubFrom(a,intloc 0); set ss = IExec(I1,s); set s0 = Initialize s; set ss0 = Initialize ss; set P = if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)); set s21 = s0 +* (loop P ';' SCM+FSA-Stop +* Start-At insloc 0); set s31 = ss0 +* (loop P ';' SCM+FSA-Stop +* Start-At insloc 0); A3: Times(a,I) = if>0(a,loop if=0(a,Goto insloc 2,I ';' SubFrom(a,intloc 0)), SCM+FSA-Stop) by Def5; A4: s0.intloc 0 = 1 & s0.a > 0 by A2,SCMFSA6C:3; then consider s2 be State of SCM+FSA, k be Nat such that A5: s2 = s0 +* (loop P +* Start-At insloc 0) and k = LifeSpan (s0 +* (P +* Start-At insloc 0)) + 1 and A6: (Computation s2).k.a = s0.a - 1 and A7: (Computation s2).k.intloc 0 = 1 and A8: (for b being read-write Int-Location st b <> a holds (Computation s2).k.b = IExec(I,s0).b) and A9: (for f being FinSeq-Location holds (Computation s2).k.f = IExec(I,s0).f) and A10: IC (Computation s2).k = insloc 0 and A11: for n being Nat st n <= k holds IC (Computation s2).n in dom loop P by A1,Th122; set C2 = Computation s2; A12: I is_halting_on s0 by SCMFSA7B:25; thus A13: now thus ss.a = Exec(SubFrom(a,intloc 0),IExec(I,s)).a by SCMFSA6C:7 .= IExec(I,s).a - IExec(I,s).intloc 0 by SCMFSA_2:91 .= IExec(I,s).a - 1 by A12,Th92 .= s0.a - 1 by A1,Th91 .= s.a - 1 by SCMFSA6C:3; end; A14: Directed loop P = loop P by SCMFSA8A:40; then A15: Directed loop P is_pseudo-closed_on s0 by A1,A4,Th117; then A16: IExec(Times(a,I),s0) | D = IExec(loop P ';' SCM+FSA-Stop,s0) | D by A3,A4,Th69; SubFrom(a,intloc 0) does_not_destroy intloc 0 by SCMFSA7B:14; then reconsider J3 = Macro SubFrom(a,intloc 0) as good Macro-Instruction by Th99; I1 = I ';' J3 by SCMFSA6A:def 6; then A17: I1 is good Macro-Instruction & I1 is_halting_on Initialize s & I1 is_closed_on Initialize s by SCMFSA7B:24,25; then A18: ss.intloc 0 = 1 by Th96; A19: now let b be Int-Location; per cases; suppose b = intloc 0; hence C2.k.b = IExec(I1,s).b by A7,A17,Th96; suppose b = a; hence C2.k.b = IExec(I1,s).b by A6,A13,SCMFSA6C:3; suppose A20: b <> a & b <> intloc 0; then reconsider bb = b as read-write Int-Location by SF_MASTR:def 5; thus C2.k.b = IExec(I,s0).bb by A8,A20 .= Exec(SubFrom(a,intloc 0),IExec(I,s0)).b by A20,SCMFSA_2:91 .= IExec(I1,s0).b by SCMFSA6C:7 .= IExec(I1,s).b by Th17; end; now let f be FinSeq-Location; thus C2.k.f = IExec(I,s0).f by A9 .= Exec(SubFrom(a,intloc 0),IExec(I,s0)).f by SCMFSA_2:91 .= IExec(I1,s0).f by SCMFSA6C:8 .= IExec(I1,s).f by Th17; end; then A21: C2.k | D = ss | D by A19,SCMFSA6A:38; A22: loop P is_pseudo-closed_on s0 by A1,A4,Th117; insloc 0 in dom P by Th54; then A23: insloc 0 in dom loop P by Th106; per cases; suppose A24: ss.a = 0; loop P c= loop P +* Start-At insloc 0 by SCMFSA8A:9; then A25: dom loop P c= dom (loop P +* Start-At insloc 0) by GRFUNC_1:8; A26: C2.k.insloc 0 = s2.insloc 0 by AMI_1:54 .= (loop P +* Start-At insloc 0).insloc 0 by A5,A23,A25,FUNCT_4:14 .= (loop P).insloc 0 by A23,SCMFSA6B:7; A27: C2.k.a = 0 by A6,A13,A24,SCMFSA6C:3; A28: C2.(k + 1) = Following C2.k by AMI_1:def 19 .= Exec(CurInstr C2.k,C2.k) by AMI_1:def 18 .= Exec(C2.k.insloc 0,C2.k) by A10,AMI_1:def 17; A29: P.insloc 0 = a =0_goto insloc (card I1 + 3) by Th55; then P.insloc 0 <> halt SCM+FSA & insloc 0 in dom P by Lm3,Th54; then A30: C2.k.insloc 0 = a =0_goto insloc (card I1 + 3) by A26,A29,Th108; then InsCode C2.k.insloc 0 = 7 by SCMFSA_2:48; then InsCode C2.k.insloc 0 in {0,6,7,8} by ENUMSET1:19; then A31: C2.k | D = C2.(k + 1) | D by A28,Th32; A32: IC C2.(k + 1) = C2.(k + 1).IC SCM+FSA by AMI_1:def 15 .= insloc (card I1 + 3) by A27,A28,A30,SCMFSA_2:96; A33: card I1 + (3 + 2) = card I1 + (4 + 1) .= card I1 + 1 + 4 by XCMPLX_1:1 .= card Goto insloc 2 + card I1 + 4 by SCMFSA8A:29 .= card P by SCMFSA8B:14 .= card dom P by PRE_CIRC:21 .= card dom loop P by Th106 .= card loop P by PRE_CIRC:21; then card I1 + 3 + 0 < card loop P by REAL_1:53; then A34: insloc (card I1 + 3) in dom loop P by SCMFSA6A:15; loop P c= loop P +* Start-At insloc 0 by SCMFSA8A:9; then A35: dom loop P c= dom (loop P +* Start-At insloc 0) by GRFUNC_1:8; A36: C2.(k + 1).insloc (card I1 + 3) = s2.insloc (card I1 + 3) by AMI_1:54 .= (loop P +* Start-At insloc 0).insloc (card I1 + 3) by A5,A34,A35,FUNCT_4 :14 .= (loop P).insloc (card I1 + 3) by A34,SCMFSA6B:7; A37: P.insloc (card I1 + 3) = goto insloc (card I1 + 5) by Th65; then P.insloc (card I1 + 3) <> halt SCM+FSA by Lm2; then A38: C2.(k + 1).insloc (card I1 + 3) = goto insloc (card I1 + 5) by A36,A37,Th108; A39: C2.(k + (1 + 1)) = C2.(k + 1 + 1) by XCMPLX_1:1 .= Following C2.(k + 1) by AMI_1:def 19 .= Exec(CurInstr C2.(k + 1),C2.(k + 1)) by AMI_1:def 18 .= Exec(C2.(k + 1).insloc (card I1 + 3),C2.(k + 1)) by A32,AMI_1:def 17; A40: IC C2.(k + 2) = C2.(k + 2).IC SCM+FSA by AMI_1:def 15 .= insloc (card I1 + 5) by A38,A39,SCMFSA_2:95 .= insloc card ProgramPart loop P by A33,AMI_5:72; now let n be Nat; assume A41: not IC (Computation s2).n in dom loop P; then k < n by A11; then k + 1 <= n by INT_1:20; then k + 1 < n by A32,A34,A41,REAL_1:def 5; then k + 1 + 1 <= n by INT_1:20; hence k + (1 + 1) <= n by XCMPLX_1:1; end; then A42: k + 2 = pseudo-LifeSpan(s0,loop P) by A5,A22,A40,SCMFSA8A:def 5; InsCode C2.(k + 1).insloc (card I1 + 3) = 6 by A38,SCMFSA_2:47; then InsCode C2.(k + 1).insloc (card I1 + 3) in {0,6,7,8} by ENUMSET1:19; then A43: C2.k | D = C2.(k + 2) | D by A31,A39,Th32; thus IExec(Times(a,I),s) | D = IExec(Times(a,I),s0) | D by Th17 .= IExec(loop P ';' SCM+FSA-Stop,s) | D by A16,Th17 .= (Result (s +* Initialized (loop P ';' SCM+FSA-Stop)) +* s | A) | D by SCMFSA6B:def 1 .= (Result s21 +* s | A) | D by SCMFSA8A:13 .= (Result s21) | D by Th35 .= IExec(I1,s) | D by A5,A14,A15,A21,A42,A43,Th59 .= IExec(Times(a,I),IExec(I1,s)) | D by A18,A24,Th123; suppose A44: ss.a <> 0; s.a >= 0 + 1 by A2,INT_1:20; then A45: ss.intloc 0 = 1 & ss.a > 0 by A13,A17,A44,Th96,REAL_1:84; then A46: Directed loop P is_pseudo-closed_on ss by A1,A14,Th117; A47: k < pseudo-LifeSpan(s0,loop P) by A5,A11,A22,Th2; then A48: (Computation s21).k | D = ss | D by A5,A14,A15,A21,Th58; A49: now A50: ss0 | D = s31 | D by SCMFSA8A:11; hereby let a be Int-Location; per cases; suppose A51: a = intloc 0; thus (Computation s21).k.a = ss.a by A48,SCMFSA6A:38 .= 1 by A51,SCMFSA6B:35 .= ss0.a by A51,SCMFSA6C:3 .= s31.a by A50,SCMFSA6A:38; suppose a <> intloc 0; then A52: a is read-write Int-Location by SF_MASTR:def 5; thus (Computation s21).k.a = ss.a by A48,SCMFSA6A:38 .= ss0.a by A52,SCMFSA6C:3 .= s31.a by A50,SCMFSA6A:38; end; let f be FinSeq-Location; thus (Computation s21).k.f = ss.f by A48,SCMFSA6A:38 .= ss0.f by SCMFSA6C:3 .= s31.f by A50,SCMFSA6A:38; end; Directed loop P = loop P by SCMFSA8A:40; then s0 | D = s21 | D & Directed loop P is_pseudo-closed_on s0 by A1,A4,Th117,SCMFSA8A:11; then Directed loop P is_pseudo-closed_on s21 by Th52; then A53: loop P ';' SCM+FSA-Stop +* Start-At insloc 0 c= s21 & loop P ';' SCM+FSA-Stop +* Start-At insloc 0 c= s31 & loop P ';' SCM+FSA-Stop is_closed_on s21 & loop P ';' SCM+FSA-Stop is_halting_on s21 by Th58,FUNCT_4:26; IC (Computation s21).k = IC C2.k by A5,A14,A15,A47,Th58 .= IC (ss0 +* (loop P ';' SCM+FSA-Stop) +* Start-At insloc 0) by A10,AMI_5: 79 .= IC s31 by FUNCT_4:15; then (Computation s21).k,s31 equal_outside A by A49,SCMFSA6A:28; then A54: Result s21,Result s31 equal_outside A by A53,Th103; IExec(loop P ';' SCM+FSA-Stop,s0) | D = IExec(loop P ';' SCM+FSA-Stop,s) | D by Th17 .= (Result (s +* Initialized (loop P ';' SCM+FSA-Stop)) +* s | A) | D by SCMFSA6B:def 1 .= (Result s21 +* s | A) | D by SCMFSA8A:13 .= (Result s21) | D by Th35 .= (Result s31) | D by A54,SCMFSA6A:39 .= (Result s31 +* ss | A) | D by Th35 .= (Result (ss +* Initialized (loop P ';' SCM+FSA-Stop)) +* ss | A) | D by SCMFSA8A:13 .= IExec(loop P ';' SCM+FSA-Stop,IExec(I1,s)) | D by SCMFSA6B:def 1 .= IExec(Times(a,I),IExec(I1,s)) | D by A3,A45,A46,Th69; hence IExec(Times(a,I),s) | D = IExec(Times(a,I),IExec(I ';' SubFrom(a,intloc 0),s)) | D by A16,Th17; end; begin ::example theorem for s being State of SCM+FSA, a,b,c being read-write Int-Location st a <> b & a <> c & b <> c & s.a >= 0 holds IExec(Times(a,Macro AddTo(b,c)),s).b = s.b + s.c * s.a proof let s be State of SCM+FSA; let a,b,c be read-write Int-Location; assume A1: a <> b & a <> c & b <> c; assume A2: s.a >= 0; set P = Times(a,Macro AddTo(b,c)); defpred P[Nat] means for s being State of SCM+FSA st s.a = $1 holds IExec(Times(a,Macro AddTo(b,c)),s).b = s.b + s.c * s.a; AddTo(b,c) does_not_destroy intloc 0 by SCMFSA7B:13; then reconsider I = Macro AddTo(b,c) as good parahalting Macro-Instruction by Th99; A3: P[0] proof let s be State of SCM+FSA; set s0 = Initialize s; assume A4: s.a = 0; then s0.intloc 0 = 1 & s0.a = 0 by SCMFSA6C:3; then A5: IExec(Times(a,I),s0) | D = s0 | D by Th123; thus IExec(P,s).b = IExec(P,s0).b by Th17 .= s0.b by A5,SCMFSA6A:38 .= s.b + s.c * s.a by A4,SCMFSA6C:3; end; A6: for k being Nat st P[k] holds P[k + 1] proof let k be Nat; assume A7: P[k]; let s be State of SCM+FSA; assume A8: s.a = k + 1; A9: IExec(I ';' SubFrom(a,intloc 0),s).b = Exec(SubFrom(a,intloc 0),IExec(I,s)).b by SCMFSA6C:7 .= IExec(I,s).b by A1,SCMFSA_2:91 .= Exec(AddTo(b,c),Initialize s).b by SCMFSA6C:6 .= (Initialize s).b + (Initialize s).c by SCMFSA_2:90 .= (Initialize s).b + s.c by SCMFSA6C:3 .= s.b + s.c by SCMFSA6C:3; A10: IExec(I ';' SubFrom(a,intloc 0),s).c = Exec(SubFrom(a,intloc 0),IExec(I,s)).c by SCMFSA6C:7 .= IExec(I,s).c by A1,SCMFSA_2:91 .= Exec(AddTo(b,c),Initialize s).c by SCMFSA6C:6 .= (Initialize s).c by A1,SCMFSA_2:90 .= s.c by SCMFSA6C:3; s.a >= 1 & 1 > 0 by A8,NAT_1:29; then A11: s.a > 0; AddTo(b,c) does_not_destroy a by A1,SCMFSA7B:13; then I does_not_destroy a by Th77; then A12: IExec(I ';' SubFrom(a,intloc 0),s).a = s.a - 1 & IExec(Times(a,I),s) | D = IExec(Times(a,I),IExec(I ';' SubFrom(a,intloc 0),s)) | D by A11,Th124 ; then IExec(I ';' SubFrom(a,intloc 0),s).a = k by A8,XCMPLX_1:26; then IExec(Times(a,I),IExec(I ';' SubFrom(a,intloc 0),s)).b = s.b + s.c + s.c * (s.a - 1) by A7,A9,A10,A12 .= s.b + s.c + (s.c * s.a - s.c * 1) by XCMPLX_1:40 .= s.b + s.c + s.c * s.a - s.c by XCMPLX_1:29 .= s.b + s.c * s.a + s.c - s.c by XCMPLX_1:1 .= s.b + s.c * s.a by XCMPLX_1:26; hence thesis by A12,SCMFSA6A:38; end; A13: for k being Nat holds P[k] from Ind(A3,A6); reconsider sa = s.a as Nat by A2,INT_1:16; P[sa] by A13; hence IExec(Times(a,Macro AddTo(b,c)),s).b = s.b + s.c * s.a; end;