:: SCMFSA8B semantic presentation begin set A = NAT ; set D = Data-Locations ; set SA0 = Start-At (0,SCM+FSA); Lm1: for I, J being Program of SCM+FSA holds Reloc (J,(card I)) c= I ";" J by FUNCT_4:25; theorem Th1: :: SCMFSA8B:1 for P being Instruction-Sequence of SCM+FSA for I being Program of SCM+FSA for s being State of SCM+FSA st I is_closed_on s,P holds 0 in dom I proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for I being Program of SCM+FSA for s being State of SCM+FSA st I is_closed_on s,P holds 0 in dom I let I be Program of SCM+FSA; ::_thesis: for s being State of SCM+FSA st I is_closed_on s,P holds 0 in dom I let s be State of SCM+FSA; ::_thesis: ( I is_closed_on s,P implies 0 in dom I ) assume A1: I is_closed_on s,P ; ::_thesis: 0 in dom I A2: IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; IC (Comput ((P +* I),(Initialize s),0)) = (Initialize s) . (IC ) .= IC (Start-At (0,SCM+FSA)) by A2, FUNCT_4:13 .= 0 by FUNCOP_1:72 ; hence 0 in dom I by A1, SCMFSA7B:def_6; ::_thesis: verum end; theorem :: SCMFSA8B:2 canceled; theorem Th3: :: SCMFSA8B:3 for P1, P2 being Instruction-Sequence of SCM+FSA for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA st DataPart s1 = DataPart s2 & I is_closed_on s1,P1 holds I is_closed_on s2,P2 proof let P1, P2 be Instruction-Sequence of SCM+FSA; ::_thesis: for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA st DataPart s1 = DataPart s2 & I is_closed_on s1,P1 holds I is_closed_on s2,P2 let s1, s2 be State of SCM+FSA; ::_thesis: for I being Program of SCM+FSA st DataPart s1 = DataPart s2 & I is_closed_on s1,P1 holds I is_closed_on s2,P2 let I be Program of SCM+FSA; ::_thesis: ( DataPart s1 = DataPart s2 & I is_closed_on s1,P1 implies I is_closed_on s2,P2 ) set S1 = Initialize s1; set S2 = Initialize s2; assume A1: DataPart s1 = DataPart s2 ; ::_thesis: ( not I is_closed_on s1,P1 or I is_closed_on s2,P2 ) A2: DataPart (Comput ((P1 +* I),(Initialize s1),0)) = DataPart s1 by MEMSTR_0:79 .= DataPart (Comput ((P2 +* I),(Initialize s2),0)) by A1, MEMSTR_0:79 ; assume A3: I is_closed_on s1,P1 ; ::_thesis: I is_closed_on s2,P2 then A4: 0 in dom I by Th1; defpred S1[ Nat] means ( IC (Comput ((P1 +* I),(Initialize s1),$1)) = IC (Comput ((P2 +* I),(Initialize s2),$1)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),$1))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),$1))) & DataPart (Comput ((P1 +* I),(Initialize s1),$1)) = DataPart (Comput ((P2 +* I),(Initialize s2),$1)) ); A5: now__::_thesis:_for_k_being_Element_of_NAT_st_S1[k]_holds_ S1[k_+_1] let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) A6: Comput ((P2 +* I),(Initialize s2),(k + 1)) = Following ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) by EXTPRO_1:3 .= Exec ((CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k)))),(Comput ((P2 +* I),(Initialize s2),k))) ; assume A7: S1[k] ; ::_thesis: S1[k + 1] then A8: for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f by SCMFSA_M:2; for a being Int-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . a = (Comput ((P2 +* I),(Initialize s2),k)) . a by A7, SCMFSA_M:2; then A9: Comput ((P1 +* I),(Initialize s1),k) = Comput ((P2 +* I),(Initialize s2),k) by A7, A8, SCMFSA_2:61; A10: IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) in dom I by A3, SCMFSA7B:def_6; Comput ((P1 +* I),(Initialize s1),(k + 1)) = Following ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) by EXTPRO_1:3 .= Exec ((CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k)))),(Comput ((P1 +* I),(Initialize s1),k))) ; then A11: Comput ((P1 +* I),(Initialize s1),(k + 1)) = Comput ((P2 +* I),(Initialize s2),(k + 1)) by A7, A9, A6; A12: IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) = IC (Comput ((P2 +* I),(Initialize s2),(k + 1))) by A11; A13: (P1 +* I) /. (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) by PBOOLE:143; A14: (P2 +* I) /. (IC (Comput ((P2 +* I),(Initialize s2),(k + 1)))) = (P2 +* I) . (IC (Comput ((P2 +* I),(Initialize s2),(k + 1)))) by PBOOLE:143; A15: I c= P1 +* I by FUNCT_4:25; A16: I c= P2 +* I by FUNCT_4:25; CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),(k + 1)))) = I . (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) by A10, A13, A15, GRFUNC_1:2 .= CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),(k + 1)))) by A12, A10, A14, A16, GRFUNC_1:2 ; hence S1[k + 1] by A11; ::_thesis: verum end; A17: (P1 +* I) /. (IC (Comput ((P1 +* I),(Initialize s1),0))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),0))) by PBOOLE:143; A18: (P2 +* I) /. (IC (Comput ((P2 +* I),(Initialize s2),0))) = (P2 +* I) . (IC (Comput ((P2 +* I),(Initialize s2),0))) by PBOOLE:143; A19: IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; then A20: IC (Comput ((P2 +* I),(Initialize s2),0)) = IC (Start-At (0,SCM+FSA)) by FUNCT_4:13 .= 0 by FUNCOP_1:72 ; A21: IC (Comput ((P1 +* I),(Initialize s1),0)) = IC (Start-At (0,SCM+FSA)) by A19, FUNCT_4:13 .= 0 by FUNCOP_1:72 ; then CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),0))) = I . 0 by A4, A17, FUNCT_4:13 .= CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),0))) by A20, A4, A18, FUNCT_4:13 ; then A22: S1[ 0 ] by A21, A20, A2; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P2_+*_I),(Initialize_s2),k))_in_dom_I let k be Element of NAT ; ::_thesis: IC (Comput ((P2 +* I),(Initialize s2),k)) in dom I A23: IC (Comput ((P1 +* I),(Initialize s1),k)) in dom I by A3, SCMFSA7B:def_6; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A22, A5); hence IC (Comput ((P2 +* I),(Initialize s2),k)) in dom I by A23; ::_thesis: verum end; hence I is_closed_on s2,P2 by SCMFSA7B:def_6; ::_thesis: verum end; theorem Th4: :: SCMFSA8B:4 for s1, s2 being State of SCM+FSA for I, J being Program of SCM+FSA st DataPart s1 = DataPart s2 holds Initialize s1 = Initialize s2 by MEMSTR_0:80; theorem Th5: :: SCMFSA8B:5 for P1, P2 being Instruction-Sequence of SCM+FSA for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA st DataPart s1 = DataPart s2 & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) proof let P1, P2 be Instruction-Sequence of SCM+FSA; ::_thesis: for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA st DataPart s1 = DataPart s2 & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) let s1, s2 be State of SCM+FSA; ::_thesis: for I being Program of SCM+FSA st DataPart s1 = DataPart s2 & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) let I be Program of SCM+FSA; ::_thesis: ( DataPart s1 = DataPart s2 & I is_closed_on s1,P1 & I is_halting_on s1,P1 implies ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) set S1 = Initialize s1; set S2 = Initialize s2; defpred S1[ Nat] means ( IC (Comput ((P1 +* I),(Initialize s1),$1)) = IC (Comput ((P2 +* I),(Initialize s2),$1)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),$1))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),$1))) & DataPart (Comput ((P1 +* I),(Initialize s1),$1)) = DataPart (Comput ((P2 +* I),(Initialize s2),$1)) ); A1: IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; A2: IC (Comput ((P1 +* I),(Initialize s1),0)) = IC (Start-At (0,SCM+FSA)) by A1, FUNCT_4:13 .= 0 by FUNCOP_1:72 ; A3: IC (Comput ((P2 +* I),(Initialize s2),0)) = IC (Start-At (0,SCM+FSA)) by A1, FUNCT_4:13 .= 0 by FUNCOP_1:72 ; assume DataPart s1 = DataPart s2 ; ::_thesis: ( not I is_closed_on s1,P1 or not I is_halting_on s1,P1 or ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) then A4: Comput ((P1 +* I),(Initialize s1),0) = Comput ((P2 +* I),(Initialize s2),0) by Th4; assume A5: I is_closed_on s1,P1 ; ::_thesis: ( not I is_halting_on s1,P1 or ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) A6: now__::_thesis:_for_k_being_Element_of_NAT_st_S1[k]_holds_ S1[k_+_1] let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) A7: Comput ((P2 +* I),(Initialize s2),(k + 1)) = Following ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) by EXTPRO_1:3 .= Exec ((CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k)))),(Comput ((P2 +* I),(Initialize s2),k))) ; assume A8: S1[k] ; ::_thesis: S1[k + 1] then A9: for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f by SCMFSA_M:2; for a being Int-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . a = (Comput ((P2 +* I),(Initialize s2),k)) . a by A8, SCMFSA_M:2; then A10: Comput ((P1 +* I),(Initialize s1),k) = Comput ((P2 +* I),(Initialize s2),k) by A8, A9, SCMFSA_2:61; A11: IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) in dom I by A5, SCMFSA7B:def_6; Comput ((P1 +* I),(Initialize s1),(k + 1)) = Following ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) by EXTPRO_1:3 .= Exec ((CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k)))),(Comput ((P1 +* I),(Initialize s1),k))) ; then A12: Comput ((P1 +* I),(Initialize s1),(k + 1)) = Comput ((P2 +* I),(Initialize s2),(k + 1)) by A8, A10, A7; A13: IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) = IC (Comput ((P2 +* I),(Initialize s2),(k + 1))) by A12; A14: (P1 +* I) /. (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) by PBOOLE:143; A15: (P2 +* I) /. (IC (Comput ((P2 +* I),(Initialize s2),(k + 1)))) = (P2 +* I) . (IC (Comput ((P2 +* I),(Initialize s2),(k + 1)))) by PBOOLE:143; A16: I c= P1 +* I by FUNCT_4:25; A17: I c= P2 +* I by FUNCT_4:25; CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),(k + 1)))) = I . (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) by A11, A14, A16, GRFUNC_1:2 .= CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),(k + 1)))) by A13, A11, A15, A17, GRFUNC_1:2 ; hence S1[k + 1] by A12; ::_thesis: verum end; assume I is_halting_on s1,P1 ; ::_thesis: ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) then P1 +* I halts_on Initialize s1 by SCMFSA7B:def_7; then consider m being Element of NAT such that A18: CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),m))) = halt SCM+FSA by EXTPRO_1:29; A19: (P1 +* I) /. (IC (Comput ((P1 +* I),(Initialize s1),0))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),0))) by PBOOLE:143; A20: (P2 +* I) /. (IC (Comput ((P2 +* I),(Initialize s2),0))) = (P2 +* I) . (IC (Comput ((P2 +* I),(Initialize s2),0))) by PBOOLE:143; A21: 0 in dom I by A5, Th1; then CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),0))) = I . 0 by A2, A19, FUNCT_4:13 .= CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),0))) by A3, A21, A20, FUNCT_4:13 ; then A22: S1[ 0 ] by A4; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P2_+*_I),(Initialize_s2),k))_in_dom_I let k be Element of NAT ; ::_thesis: IC (Comput ((P2 +* I),(Initialize s2),k)) in dom I A23: IC (Comput ((P1 +* I),(Initialize s1),k)) in dom I by A5, SCMFSA7B:def_6; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A22, A6); hence IC (Comput ((P2 +* I),(Initialize s2),k)) in dom I by A23; ::_thesis: verum end; hence I is_closed_on s2,P2 by SCMFSA7B:def_6; ::_thesis: I is_halting_on s2,P2 for k being Element of NAT holds S1[k] from NAT_1:sch_1(A22, A6); then CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),m))) = halt SCM+FSA by A18; then P2 +* I halts_on Initialize s2 by EXTPRO_1:29; hence I is_halting_on s2,P2 by SCMFSA7B:def_7; ::_thesis: verum end; theorem Th6: :: SCMFSA8B:6 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA holds ( I is_closed_on Initialized s,P iff I is_closed_on s +* (Initialize ((intloc 0) .--> 1)),P +* J ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA holds ( I is_closed_on Initialized s,P iff I is_closed_on s +* (Initialize ((intloc 0) .--> 1)),P +* J ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA holds ( I is_closed_on Initialized s,P iff I is_closed_on s +* (Initialize ((intloc 0) .--> 1)),P +* J ) let I, J be Program of SCM+FSA; ::_thesis: ( I is_closed_on Initialized s,P iff I is_closed_on s +* (Initialize ((intloc 0) .--> 1)),P +* J ) DataPart (Initialized s) = DataPart (s +* (Initialize ((intloc 0) .--> 1))) ; hence ( I is_closed_on Initialized s,P iff I is_closed_on s +* (Initialize ((intloc 0) .--> 1)),P +* J ) by Th3; ::_thesis: verum end; theorem Th7: :: SCMFSA8B:7 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for l being Element of NAT holds ( I is_closed_on s,P iff I is_closed_on s +* (Start-At (0,SCM+FSA)),P +* I ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for l being Element of NAT holds ( I is_closed_on s,P iff I is_closed_on s +* (Start-At (0,SCM+FSA)),P +* I ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for l being Element of NAT holds ( I is_closed_on s,P iff I is_closed_on s +* (Start-At (0,SCM+FSA)),P +* I ) let I, J be Program of SCM+FSA; ::_thesis: for l being Element of NAT holds ( I is_closed_on s,P iff I is_closed_on s +* (Start-At (0,SCM+FSA)),P +* I ) let l be Element of NAT ; ::_thesis: ( I is_closed_on s,P iff I is_closed_on s +* (Start-At (0,SCM+FSA)),P +* I ) DataPart s = DataPart (Initialize s) by MEMSTR_0:79; hence ( I is_closed_on s,P iff I is_closed_on s +* (Start-At (0,SCM+FSA)),P +* I ) by Th3; ::_thesis: verum end; theorem Th8: :: SCMFSA8B:8 for P1, P2 being Instruction-Sequence of SCM+FSA for s1 being 0 -started State of SCM+FSA for s2 being State of SCM+FSA for I being Program of SCM+FSA st I is_closed_on s1,P1 & I c= P1 holds for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) proof let P1, P2 be Instruction-Sequence of SCM+FSA; ::_thesis: for s1 being 0 -started State of SCM+FSA for s2 being State of SCM+FSA for I being Program of SCM+FSA st I is_closed_on s1,P1 & I c= P1 holds for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) let s1 be 0 -started State of SCM+FSA; ::_thesis: for s2 being State of SCM+FSA for I being Program of SCM+FSA st I is_closed_on s1,P1 & I c= P1 holds for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) let s2 be State of SCM+FSA; ::_thesis: for I being Program of SCM+FSA st I is_closed_on s1,P1 & I c= P1 holds for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) let I be Program of SCM+FSA; ::_thesis: ( I is_closed_on s1,P1 & I c= P1 implies for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) ) A1: Start-At (0,SCM+FSA) c= s1 by MEMSTR_0:29; assume A2: I is_closed_on s1,P1 ; ::_thesis: ( not I c= P1 or for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) ) assume A3: I c= P1 ; ::_thesis: for n being Element of NAT st IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 holds for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) let n be Element of NAT ; ::_thesis: ( IC s2 = n & DataPart s1 = DataPart s2 & Reloc (I,n) c= P2 implies for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) ) A4: IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; defpred S1[ Nat] means ( (IC (Comput (P1,s1,$1))) + n = IC (Comput (P2,s2,$1)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,$1)))),n) = CurInstr (P2,(Comput (P2,s2,$1))) & DataPart (Comput (P1,s1,$1)) = DataPart (Comput (P2,s2,$1)) ); A5: IC (Comput (P1,s1,0)) = IC s1 .= IC (Start-At (0,SCM+FSA)) by A1, A4, GRFUNC_1:2 .= 0 by FUNCOP_1:72 ; assume A6: IC s2 = n ; ::_thesis: ( not DataPart s1 = DataPart s2 or not Reloc (I,n) c= P2 or for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) ) A7: 0 in dom I by A2, Th1; then A8: 0 + n in dom (Reloc (I,n)) by COMPOS_1:46; IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; then A9: P1 . (IC s1) = P1 . (IC (Start-At (0,SCM+FSA))) by A1, GRFUNC_1:2 .= P1 . 0 by FUNCOP_1:72 .= I . 0 by A7, A3, GRFUNC_1:2 ; assume DataPart s1 = DataPart s2 ; ::_thesis: ( not Reloc (I,n) c= P2 or for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) ) then A10: DataPart (Comput (P1,s1,0)) = DataPart s2 .= DataPart (Comput (P2,s2,0)) ; assume A11: Reloc (I,n) c= P2 ; ::_thesis: for i being Element of NAT holds ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) let i be Element of NAT ; ::_thesis: ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) A12: P2 /. (IC s2) = P2 . (IC s2) by PBOOLE:143; A13: CurInstr (P1,s1) = I . 0 by A9, PBOOLE:143; IncAddr ((CurInstr (P1,(Comput (P1,s1,0)))),n) = IncAddr ((CurInstr (P1,s1)),n) .= (Reloc (I,n)) . (0 + n) by A13, A7, COMPOS_1:35 .= CurInstr (P2,s2) by A6, A8, A12, A11, GRFUNC_1:2 .= CurInstr (P2,(Comput (P2,s2,0))) ; then A14: S1[ 0 ] by A6, A5, A10; A15: for k being Element of NAT st S1[k] holds S1[k + 1] proof let k be Element of NAT ; ::_thesis: ( S1[k] implies S1[k + 1] ) A16: Comput (P1,s1,(k + 1)) = Following (P1,(Comput (P1,s1,k))) by EXTPRO_1:3 .= Exec ((CurInstr (P1,(Comput (P1,s1,k)))),(Comput (P1,s1,k))) ; reconsider l = IC (Comput (P1,s1,(k + 1))) as Element of NAT ; reconsider j = CurInstr (P1,(Comput (P1,s1,(k + 1)))) as Instruction of SCM+FSA ; A17: Comput (P2,s2,(k + 1)) = Following (P2,(Comput (P2,s2,k))) by EXTPRO_1:3 .= Exec ((CurInstr (P2,(Comput (P2,s2,k)))),(Comput (P2,s2,k))) ; A18: Initialize s1 = s1 by A1, FUNCT_4:98; A19: P1 = P1 +* I by A3, FUNCT_4:98; then A20: IC (Comput (P1,s1,(k + 1))) in dom I by A2, A18, SCMFSA7B:def_6; assume A21: S1[k] ; ::_thesis: S1[k + 1] hence A22: (IC (Comput (P1,s1,(k + 1)))) + n = IC (Comput (P2,s2,(k + 1))) by A16, A17, SCMFSA6A:8; ::_thesis: ( IncAddr ((CurInstr (P1,(Comput (P1,s1,(k + 1))))),n) = CurInstr (P2,(Comput (P2,s2,(k + 1)))) & DataPart (Comput (P1,s1,(k + 1))) = DataPart (Comput (P2,s2,(k + 1))) ) then A23: IC (Comput (P2,s2,(k + 1))) in dom (Reloc (I,n)) by A20, COMPOS_1:46; A24: l in dom I by A18, A2, A19, SCMFSA7B:def_6; j = P1 . (IC (Comput (P1,s1,(k + 1)))) by PBOOLE:143 .= I . l by A20, A3, GRFUNC_1:2 ; hence IncAddr ((CurInstr (P1,(Comput (P1,s1,(k + 1))))),n) = (Reloc (I,n)) . (l + n) by A24, COMPOS_1:35 .= P2 . (IC (Comput (P2,s2,(k + 1)))) by A23, A22, A11, GRFUNC_1:2 .= CurInstr (P2,(Comput (P2,s2,(k + 1)))) by PBOOLE:143 ; ::_thesis: DataPart (Comput (P1,s1,(k + 1))) = DataPart (Comput (P2,s2,(k + 1))) thus DataPart (Comput (P1,s1,(k + 1))) = DataPart (Comput (P2,s2,(k + 1))) by A21, A16, A17, SCMFSA6A:8; ::_thesis: verum end; for k being Element of NAT holds S1[k] from NAT_1:sch_1(A14, A15); hence ( (IC (Comput (P1,s1,i))) + n = IC (Comput (P2,s2,i)) & IncAddr ((CurInstr (P1,(Comput (P1,s1,i)))),n) = CurInstr (P2,(Comput (P2,s2,i))) & DataPart (Comput (P1,s1,i)) = DataPart (Comput (P2,s2,i)) ) ; ::_thesis: verum end; theorem Th9: :: SCMFSA8B:9 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for i being parahalting keeping_0 Instruction of SCM+FSA for J being parahalting Program of SCM+FSA for a being Int-Location holds (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(Exec (i,(Initialized s))))) . a proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for i being parahalting keeping_0 Instruction of SCM+FSA for J being parahalting Program of SCM+FSA for a being Int-Location holds (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(Exec (i,(Initialized s))))) . a let s be State of SCM+FSA; ::_thesis: for i being parahalting keeping_0 Instruction of SCM+FSA for J being parahalting Program of SCM+FSA for a being Int-Location holds (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(Exec (i,(Initialized s))))) . a let i be parahalting keeping_0 Instruction of SCM+FSA; ::_thesis: for J being parahalting Program of SCM+FSA for a being Int-Location holds (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(Exec (i,(Initialized s))))) . a let J be parahalting Program of SCM+FSA; ::_thesis: for a being Int-Location holds (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(Exec (i,(Initialized s))))) . a let a be Int-Location; ::_thesis: (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(Exec (i,(Initialized s))))) . a A1: IExec ((Macro i),P,s) = Exec (i,(Initialized s)) by SCMFSA6C:5; thus (IExec ((i ";" J),P,s)) . a = (IExec (J,P,(IExec ((Macro i),P,s)))) . a by SCMFSA6C:1 .= (IExec (J,P,(IExec ((Macro i),P,s)))) . a .= (IExec (J,P,(Exec (i,(Initialized s))))) . a by A1 .= (IExec (J,P,(Exec (i,(Initialized s))))) . a ; ::_thesis: verum end; theorem Th10: :: SCMFSA8B:10 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for i being parahalting keeping_0 Instruction of SCM+FSA for J being parahalting Program of SCM+FSA for f being FinSeq-Location holds (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(Exec (i,(Initialized s))))) . f proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for i being parahalting keeping_0 Instruction of SCM+FSA for J being parahalting Program of SCM+FSA for f being FinSeq-Location holds (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(Exec (i,(Initialized s))))) . f let s be State of SCM+FSA; ::_thesis: for i being parahalting keeping_0 Instruction of SCM+FSA for J being parahalting Program of SCM+FSA for f being FinSeq-Location holds (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(Exec (i,(Initialized s))))) . f let i be parahalting keeping_0 Instruction of SCM+FSA; ::_thesis: for J being parahalting Program of SCM+FSA for f being FinSeq-Location holds (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(Exec (i,(Initialized s))))) . f let J be parahalting Program of SCM+FSA; ::_thesis: for f being FinSeq-Location holds (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(Exec (i,(Initialized s))))) . f let f be FinSeq-Location ; ::_thesis: (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(Exec (i,(Initialized s))))) . f A1: IExec ((Macro i),P,s) = Exec (i,(Initialized s)) by SCMFSA6C:5; thus (IExec ((i ";" J),P,s)) . f = (IExec (J,P,(IExec ((Macro i),P,s)))) . f by SCMFSA6C:2 .= (IExec (J,P,(IExec ((Macro i),P,s)))) . f .= (IExec (J,P,(Exec (i,(Initialized s))))) . f by A1 .= (IExec (J,P,(Exec (i,(Initialized s))))) . f ; ::_thesis: verum end; definition let a be Int-Location; let I, J be Program of SCM+FSA; func if=0 (a,I,J) -> Program of SCM+FSA equals :: SCMFSA8B:def 1 ((((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); coherence ((((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is Program of SCM+FSA ; func if>0 (a,I,J) -> Program of SCM+FSA equals :: SCMFSA8B:def 2 ((((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); coherence ((((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is Program of SCM+FSA ; end; :: deftheorem defines if=0 SCMFSA8B:def_1_:_ for a being Int-Location for I, J being Program of SCM+FSA holds if=0 (a,I,J) = ((((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); :: deftheorem defines if>0 SCMFSA8B:def_2_:_ for a being Int-Location for I, J being Program of SCM+FSA holds if>0 (a,I,J) = ((((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); definition let a be Int-Location; let I, J be Program of SCM+FSA; func if<0 (a,I,J) -> Program of SCM+FSA equals :: SCMFSA8B:def 3 if=0 (a,J,(if>0 (a,J,I))); coherence if=0 (a,J,(if>0 (a,J,I))) is Program of SCM+FSA ; end; :: deftheorem defines if<0 SCMFSA8B:def_3_:_ for a being Int-Location for I, J being Program of SCM+FSA holds if<0 (a,I,J) = if=0 (a,J,(if>0 (a,J,I))); Lm2: for a being Int-Location for I, J being Program of SCM+FSA holds ( 0 in dom (if=0 (a,I,J)) & 1 in dom (if=0 (a,I,J)) & 0 in dom (if>0 (a,I,J)) & 1 in dom (if>0 (a,I,J)) ) proof let a be Int-Location; ::_thesis: for I, J being Program of SCM+FSA holds ( 0 in dom (if=0 (a,I,J)) & 1 in dom (if=0 (a,I,J)) & 0 in dom (if>0 (a,I,J)) & 1 in dom (if>0 (a,I,J)) ) let I, J be Program of SCM+FSA; ::_thesis: ( 0 in dom (if=0 (a,I,J)) & 1 in dom (if=0 (a,I,J)) & 0 in dom (if>0 (a,I,J)) & 1 in dom (if>0 (a,I,J)) ) set i = a =0_goto ((card J) + 3); if=0 (a,I,J) = (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a =0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a =0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (Macro (a =0_goto ((card J) + 3))) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) ; then A1: dom (Macro (a =0_goto ((card J) + 3))) c= dom (if=0 (a,I,J)) by SCMFSA6A:17; A2: dom (Macro (a =0_goto ((card J) + 3))) = {0,1} by COMPOS_1:61; then A3: 1 in dom (Macro (a =0_goto ((card J) + 3))) by TARSKI:def_2; 0 in dom (Macro (a =0_goto ((card J) + 3))) by A2, TARSKI:def_2; hence ( 0 in dom (if=0 (a,I,J)) & 1 in dom (if=0 (a,I,J)) ) by A1, A3; ::_thesis: ( 0 in dom (if>0 (a,I,J)) & 1 in dom (if>0 (a,I,J)) ) set i = a >0_goto ((card J) + 3); if>0 (a,I,J) = (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a >0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a >0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (Macro (a >0_goto ((card J) + 3))) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) ; then A4: dom (Macro (a >0_goto ((card J) + 3))) c= dom (if>0 (a,I,J)) by SCMFSA6A:17; A5: dom (Macro (a >0_goto ((card J) + 3))) = {0,1} by COMPOS_1:61; then A6: 1 in dom (Macro (a >0_goto ((card J) + 3))) by TARSKI:def_2; 0 in dom (Macro (a >0_goto ((card J) + 3))) by A5, TARSKI:def_2; hence ( 0 in dom (if>0 (a,I,J)) & 1 in dom (if>0 (a,I,J)) ) by A4, A6; ::_thesis: verum end; Lm3: for a being Int-Location for I, J being Program of SCM+FSA holds ( (if=0 (a,I,J)) . 0 = a =0_goto ((card J) + 3) & (if=0 (a,I,J)) . 1 = goto 2 & (if>0 (a,I,J)) . 0 = a >0_goto ((card J) + 3) & (if>0 (a,I,J)) . 1 = goto 2 ) proof let a be Int-Location; ::_thesis: for I, J being Program of SCM+FSA holds ( (if=0 (a,I,J)) . 0 = a =0_goto ((card J) + 3) & (if=0 (a,I,J)) . 1 = goto 2 & (if>0 (a,I,J)) . 0 = a >0_goto ((card J) + 3) & (if>0 (a,I,J)) . 1 = goto 2 ) let I, J be Program of SCM+FSA; ::_thesis: ( (if=0 (a,I,J)) . 0 = a =0_goto ((card J) + 3) & (if=0 (a,I,J)) . 1 = goto 2 & (if>0 (a,I,J)) . 0 = a >0_goto ((card J) + 3) & (if>0 (a,I,J)) . 1 = goto 2 ) set i = a =0_goto ((card J) + 3); A1: if=0 (a,I,J) = (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a =0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a =0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (Macro (a =0_goto ((card J) + 3))) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) ; A2: dom (Macro (a =0_goto ((card J) + 3))) = {0,1} by COMPOS_1:61; then 0 in dom (Macro (a =0_goto ((card J) + 3))) by TARSKI:def_2; hence (if=0 (a,I,J)) . 0 = (Directed (Macro (a =0_goto ((card J) + 3)))) . 0 by A1, SCMFSA8A:14 .= a =0_goto ((card J) + 3) by SCMFSA7B:1 ; ::_thesis: ( (if=0 (a,I,J)) . 1 = goto 2 & (if>0 (a,I,J)) . 0 = a >0_goto ((card J) + 3) & (if>0 (a,I,J)) . 1 = goto 2 ) 1 in dom (Macro (a =0_goto ((card J) + 3))) by A2, TARSKI:def_2; hence (if=0 (a,I,J)) . 1 = (Directed (Macro (a =0_goto ((card J) + 3)))) . 1 by A1, SCMFSA8A:14 .= goto 2 by SCMFSA7B:2 ; ::_thesis: ( (if>0 (a,I,J)) . 0 = a >0_goto ((card J) + 3) & (if>0 (a,I,J)) . 1 = goto 2 ) set i = a >0_goto ((card J) + 3); A3: if>0 (a,I,J) = (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a >0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a >0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (Macro (a >0_goto ((card J) + 3))) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) ; A4: dom (Macro (a >0_goto ((card J) + 3))) = {0,1} by COMPOS_1:61; then 0 in dom (Macro (a >0_goto ((card J) + 3))) by TARSKI:def_2; hence (if>0 (a,I,J)) . 0 = (Directed (Macro (a >0_goto ((card J) + 3)))) . 0 by A3, SCMFSA8A:14 .= a >0_goto ((card J) + 3) by SCMFSA7B:1 ; ::_thesis: (if>0 (a,I,J)) . 1 = goto 2 1 in dom (Macro (a >0_goto ((card J) + 3))) by A4, TARSKI:def_2; hence (if>0 (a,I,J)) . 1 = (Directed (Macro (a >0_goto ((card J) + 3)))) . 1 by A3, SCMFSA8A:14 .= goto 2 by SCMFSA7B:2 ; ::_thesis: verum end; theorem Th11: :: SCMFSA8B:11 for I, J being Program of SCM+FSA for a being Int-Location holds card (if=0 (a,I,J)) = ((card I) + (card J)) + 4 proof let I, J be Program of SCM+FSA; ::_thesis: for a being Int-Location holds card (if=0 (a,I,J)) = ((card I) + (card J)) + 4 let a be Int-Location; ::_thesis: card (if=0 (a,I,J)) = ((card I) + (card J)) + 4 A1: card (Stop SCM+FSA) = 1 by COMPOS_1:4; thus card (if=0 (a,I,J)) = (card ((((Macro (a =0_goto ((card J) + 3))) ";" J) ";" (Goto ((card I) + 1))) ";" I)) + 1 by A1, SCMFSA6A:21 .= ((card (((Macro (a =0_goto ((card J) + 3))) ";" J) ";" (Goto ((card I) + 1)))) + (card I)) + 1 by SCMFSA6A:21 .= (((card ((Macro (a =0_goto ((card J) + 3))) ";" J)) + (card (Goto ((card I) + 1)))) + (card I)) + 1 by SCMFSA6A:21 .= (((card ((Macro (a =0_goto ((card J) + 3))) ";" J)) + 1) + (card I)) + 1 by SCMFSA8A:15 .= ((((card (Macro (a =0_goto ((card J) + 3)))) + (card J)) + 1) + (card I)) + 1 by SCMFSA6A:21 .= (((2 + (card J)) + 1) + (card I)) + 1 by COMPOS_1:56 .= ((card I) + (card J)) + 4 ; ::_thesis: verum end; theorem Th12: :: SCMFSA8B:12 for I, J being Program of SCM+FSA for a being Int-Location holds card (if>0 (a,I,J)) = ((card I) + (card J)) + 4 proof let I, J be Program of SCM+FSA; ::_thesis: for a being Int-Location holds card (if>0 (a,I,J)) = ((card I) + (card J)) + 4 let a be Int-Location; ::_thesis: card (if>0 (a,I,J)) = ((card I) + (card J)) + 4 A1: card (Stop SCM+FSA) = 1 by COMPOS_1:4; thus card (if>0 (a,I,J)) = (card ((((Macro (a >0_goto ((card J) + 3))) ";" J) ";" (Goto ((card I) + 1))) ";" I)) + 1 by A1, SCMFSA6A:21 .= ((card (((Macro (a >0_goto ((card J) + 3))) ";" J) ";" (Goto ((card I) + 1)))) + (card I)) + 1 by SCMFSA6A:21 .= (((card ((Macro (a >0_goto ((card J) + 3))) ";" J)) + (card (Goto ((card I) + 1)))) + (card I)) + 1 by SCMFSA6A:21 .= (((card ((Macro (a >0_goto ((card J) + 3))) ";" J)) + 1) + (card I)) + 1 by SCMFSA8A:15 .= ((((card (Macro (a >0_goto ((card J) + 3)))) + (card J)) + 1) + (card I)) + 1 by SCMFSA6A:21 .= (((2 + (card J)) + 1) + (card I)) + 1 by COMPOS_1:56 .= ((card I) + (card J)) + 4 ; ::_thesis: verum end; theorem Th13: :: SCMFSA8B:13 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a = 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) let a be read-write Int-Location; ::_thesis: ( s . a = 0 & I is_closed_on s,P & I is_halting_on s,P implies ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) ) set I1 = I ";" (Stop SCM+FSA); set s1 = Initialize s; set P1 = P +* (I ";" (Stop SCM+FSA)); set s3 = Initialize s; set P3 = P +* (if=0 (a,I,J)); set s4 = Comput ((P +* (if=0 (a,I,J))),(Initialize s),1); set i = a =0_goto ((card J) + 3); A1: not a in dom (Start-At (0,SCM+FSA)) by SCMFSA_2:102; A2: 0 in dom (if=0 (a,I,J)) by Lm2; A3: (P +* (if=0 (a,I,J))) . 0 = (if=0 (a,I,J)) . 0 by A2, FUNCT_4:13 .= a =0_goto ((card J) + 3) by Lm3 ; IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; then A4: IC (Initialize s) = IC (Start-At (0,SCM+FSA)) by FUNCT_4:13 .= 0 by FUNCOP_1:72 ; A5: if=0 (a,I,J) c= P +* (if=0 (a,I,J)) by FUNCT_4:25; A6: if=0 (a,I,J) = (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25; card (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) = (card ((Macro (a =0_goto ((card J) + 3))) ";" J)) + (card (Goto ((card I) + 1))) by SCMFSA6A:21 .= (card ((Macro (a =0_goto ((card J) + 3))) ";" J)) + 1 by SCMFSA8A:15 .= ((card (Macro (a =0_goto ((card J) + 3)))) + (card J)) + 1 by SCMFSA6A:21 .= ((card J) + 2) + 1 by COMPOS_1:56 .= (card J) + (2 + 1) ; then A7: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= if=0 (a,I,J) by A6, Lm1; A8: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= P +* (if=0 (a,I,J)) by A7, A5, XBOOLE_1:1; A9: Comput ((P +* (if=0 (a,I,J))),(Initialize s),(0 + 1)) = Following ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),0))) by EXTPRO_1:3 .= Following ((P +* (if=0 (a,I,J))),(Initialize s)) .= Exec ((a =0_goto ((card J) + 3)),(Initialize s)) by A4, A3, PBOOLE:143 ; A10: for f being FinSeq-Location holds (Initialize s) . f = (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) . f by A9, SCMFSA_2:70; for a being Int-Location holds (Initialize s) . a = (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) . a by A9, SCMFSA_2:70; then A11: DataPart (Initialize s) = DataPart (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) by A10, SCMFSA_M:2; assume s . a = 0 ; ::_thesis: ( not I is_closed_on s,P or not I is_halting_on s,P or ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) ) then (Initialize s) . a = 0 by A1, FUNCT_4:11; then A12: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) = (card J) + 3 by A9, SCMFSA_2:70; assume A13: I is_closed_on s,P ; ::_thesis: ( not I is_halting_on s,P or ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) ) assume A14: I is_halting_on s,P ; ::_thesis: ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) then A15: I ";" (Stop SCM+FSA) is_closed_on s,P by A13, SCMFSA8A:30; I ";" (Stop SCM+FSA) is_halting_on s,P by A13, A14, SCMFSA8A:30; then A16: P +* (I ";" (Stop SCM+FSA)) halts_on Initialize s by SCMFSA7B:def_7; DataPart s = DataPart (Initialize s) by MEMSTR_0:79; then A17: I ";" (Stop SCM+FSA) is_closed_on Initialize s,P +* (I ";" (Stop SCM+FSA)) by A15, Th3; A18: I ";" (Stop SCM+FSA) c= P +* (I ";" (Stop SCM+FSA)) by FUNCT_4:25; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P_+*_(if=0_(a,I,J))),(Initialize_s),k))_in_dom_(if=0_(a,I,J)) let k be Element of NAT ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) percases ( 0 < k or k = 0 ) ; suppose 0 < k ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) then consider k1 being Nat such that A19: k1 + 1 = k by NAT_1:6; reconsider k1 = k1 as Element of NAT by ORDINAL1:def_12; reconsider m = IC (Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialize s),k1)) as Element of NAT ; m in dom (I ";" (Stop SCM+FSA)) by A15, SCMFSA7B:def_6; then A20: m < card (I ";" (Stop SCM+FSA)) by AFINSQ_1:66; card (Stop SCM+FSA) = 1 by COMPOS_1:4; then A21: card (I ";" (Stop SCM+FSA)) = (card I) + 1 by SCMFSA6A:21; card (if=0 (a,I,J)) = ((card I) + (card J)) + 4 by Th11 .= ((card J) + 3) + (card (I ";" (Stop SCM+FSA))) by A21 ; then A22: m + ((card J) + 3) < card (if=0 (a,I,J)) by A20, XREAL_1:6; IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) = IC (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)),k1)) by A19, EXTPRO_1:4 .= m + ((card J) + 3) by A17, A12, A11, Th8, A18, A8 ; hence IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) in dom (if=0 (a,I,J)) by A22, AFINSQ_1:66; ::_thesis: verum end; suppose k = 0 ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) hence IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) in dom (if=0 (a,I,J)) by A2, A4, EXTPRO_1:2; ::_thesis: verum end; end; end; hence if=0 (a,I,J) is_closed_on s,P by SCMFSA7B:def_6; ::_thesis: if=0 (a,I,J) is_halting_on s,P CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialize s))) + 1)))) = CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialize s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialize s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialize s))))))),((card J) + 3)) by A17, A12, A11, Th8, A8, A18 .= IncAddr ((halt SCM+FSA),((card J) + 3)) by A16, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then P +* (if=0 (a,I,J)) halts_on Initialize s by EXTPRO_1:29; hence if=0 (a,I,J) is_halting_on s,P by SCMFSA7B:def_7; ::_thesis: verum end; theorem Th14: :: SCMFSA8B:14 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a = 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let a be read-write Int-Location; ::_thesis: ( s . a = 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) set I1 = I ";" (Stop SCM+FSA); set s1 = Initialized s; set P1 = P +* (I ";" (Stop SCM+FSA)); set P3 = P +* (if=0 (a,I,J)); A1: I ";" (Stop SCM+FSA) c= P +* (I ";" (Stop SCM+FSA)) by FUNCT_4:25; set s4 = Comput ((P +* (if=0 (a,I,J))),(Initialized s),1); set i = a =0_goto ((card J) + 3); A2: if=0 (a,I,J) = (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25; A3: 0 in dom (if=0 (a,I,J)) by Lm2; A4: (P +* (if=0 (a,I,J))) . 0 = (if=0 (a,I,J)) . 0 by A3, FUNCT_4:13 .= a =0_goto ((card J) + 3) by Lm3 ; A5: dom (Initialize ((intloc 0) .--> 1)) = {(intloc 0),(IC )} by SCMFSA_M:11; ( a <> intloc 0 & a <> IC ) by SCMFSA_2:56; then A6: not a in dom (Initialize ((intloc 0) .--> 1)) by A5, TARSKI:def_2; IC in dom (Initialize ((intloc 0) .--> 1)) by MEMSTR_0:48; then A7: IC (Initialized s) = IC (Initialize ((intloc 0) .--> 1)) by FUNCT_4:13 .= 0 by MEMSTR_0:def_11 ; A8: Comput ((P +* (if=0 (a,I,J))),(Initialized s),(0 + 1)) = Following ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),0))) by EXTPRO_1:3 .= Following ((P +* (if=0 (a,I,J))),(Initialized s)) .= Exec ((a =0_goto ((card J) + 3)),(Initialized s)) by A7, A4, PBOOLE:143 ; A9: if=0 (a,I,J) c= P +* (if=0 (a,I,J)) by FUNCT_4:25; assume s . a = 0 ; ::_thesis: ( not I is_closed_on Initialized s,P or not I is_halting_on Initialized s,P or IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) then (Initialized s) . a = 0 by A6, FUNCT_4:11; then A10: IC (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) = (card J) + 3 by A8, SCMFSA_2:70; A11: for f being FinSeq-Location holds (Initialized s) . f = (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) . f by A8, SCMFSA_2:70; for a being Int-Location holds (Initialized s) . a = (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) . a by A8, SCMFSA_2:70; then A12: DataPart (Initialized s) = DataPart (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) by A11, SCMFSA_M:2; card (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) = (card ((Macro (a =0_goto ((card J) + 3))) ";" J)) + (card (Goto ((card I) + 1))) by SCMFSA6A:21 .= (card ((Macro (a =0_goto ((card J) + 3))) ";" J)) + 1 by SCMFSA8A:15 .= ((card (Macro (a =0_goto ((card J) + 3)))) + (card J)) + 1 by SCMFSA6A:21 .= ((card J) + 2) + 1 by COMPOS_1:56 .= (card J) + (2 + 1) ; then A13: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= if=0 (a,I,J) by A2, Lm1; A14: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= P +* (if=0 (a,I,J)) by A13, A9, XBOOLE_1:1; assume A15: I is_closed_on Initialized s,P ; ::_thesis: ( not I is_halting_on Initialized s,P or IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) assume A16: I is_halting_on Initialized s,P ; ::_thesis: IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) then A17: P +* (I ";" (Stop SCM+FSA)) halts_on Initialized s by A15, SCMFSA8A:34; I ";" (Stop SCM+FSA) is_closed_on Initialized s,P by A15, A16, SCMFSA8A:30; then A18: I ";" (Stop SCM+FSA) is_closed_on Initialized s,P +* (I ";" (Stop SCM+FSA)) by Th6; A19: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1)))) = CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))))),((card J) + 3)) by A18, A10, A12, Th8, A14, A1 .= IncAddr ((halt SCM+FSA),((card J) + 3)) by A17, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then A20: P +* (if=0 (a,I,J)) halts_on Initialized s by EXTPRO_1:29; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_(LifeSpan_((P_+*_(I_";"_(Stop_SCM+FSA))),(Initialized_s)))_+_1_holds_ CurInstr_((P_+*_(if=0_(a,I,J))),(Comput_((P_+*_(if=0_(a,I,J))),(Initialized_s),l)))_<>_halt_SCM+FSA let l be Element of NAT ; ::_thesis: ( l < (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 implies CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA ) assume A21: l < (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 ; ::_thesis: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA A22: Comput ((P +* (if=0 (a,I,J))),(Initialized s),0) = Initialized s ; percases ( l = 0 or l <> 0 ) ; suppose l = 0 ; ::_thesis: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA hence CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) <> halt SCM+FSA by A7, A4, A22, PBOOLE:143; ::_thesis: verum end; suppose l <> 0 ; ::_thesis: not CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) = halt SCM+FSA then consider n being Nat such that A23: l = n + 1 by NAT_1:6; assume A24: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA ; ::_thesis: contradiction reconsider n = n as Element of NAT by ORDINAL1:def_12; A25: InsCode (halt SCM+FSA) = 0 by COMPOS_1:70; InsCode (CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),n)))) = InsCode (IncAddr ((CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),n)))),((card J) + 3))) by COMPOS_0:def_9 .= InsCode (CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)),n)))) by A18, A10, A12, Th8, A14, A1 .= 0 by A23, A24, EXTPRO_1:4, A25 ; then A26: CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),n))) = halt SCM+FSA by SCMFSA_2:95; n < LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)) by A21, A23, XREAL_1:6; hence contradiction by A17, A26, EXTPRO_1:def_15; ::_thesis: verum end; end; end; then for l being Element of NAT st CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA holds (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 <= l ; then A27: LifeSpan ((P +* (if=0 (a,I,J))),(Initialized s)) = (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 by A19, A20, EXTPRO_1:def_15; A28: DataPart (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) = DataPart (Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))) by A15, A16, EXTPRO_1:23, SCMFSA8A:34 .= DataPart (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))) by A18, A10, A12, Th8, A1, A14 .= DataPart (Comput ((P +* (if=0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1))) by EXTPRO_1:4 .= DataPart (Result ((P +* (if=0 (a,I,J))),(Initialized s))) by A20, A27, EXTPRO_1:23 ; A29: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((if=0_(a,I,J)),P,s))_holds_ (IExec_((if=0_(a,I,J)),P,s))_._x_=_((IExec_((I_";"_(Stop_SCM+FSA)),P,s))_+*_(Start-At_((((card_I)_+_(card_J))_+_3),SCM+FSA)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((if=0 (a,I,J)),P,s)) implies (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 ) A30: IExec ((I ";" (Stop SCM+FSA)),P,s) = Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)) by SCMFSA6B:def_1; A31: dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) = {(IC )} by FUNCOP_1:13; A32: IExec ((if=0 (a,I,J)),P,s) = Result ((P +* (if=0 (a,I,J))),(Initialized s)) by SCMFSA6B:def_1; assume A33: x in dom (IExec ((if=0 (a,I,J)),P,s)) ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 percases ( x is Int-Location or x is FinSeq-Location or x = IC ) by A33, SCMFSA_M:1; supposeA34: x is Int-Location ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:56; then A35: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A31, TARSKI:def_1; thus (IExec ((if=0 (a,I,J)),P,s)) . x = (Result ((P +* (if=0 (a,I,J))),(Initialized s))) . x by A32 .= (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) . x by A28, A34, SCMFSA_M:2 .= (IExec ((I ";" (Stop SCM+FSA)),P,s)) . x by A30 .= ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A35, FUNCT_4:11 ; ::_thesis: verum end; supposeA36: x is FinSeq-Location ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:57; then A37: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A31, TARSKI:def_1; thus (IExec ((if=0 (a,I,J)),P,s)) . x = (Result ((P +* (if=0 (a,I,J))),(Initialized s))) . x by A32 .= (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) . x by A28, A36, SCMFSA_M:2 .= (IExec ((I ";" (Stop SCM+FSA)),P,s)) . x by A30 .= ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A37, FUNCT_4:11 ; ::_thesis: verum end; supposeA38: x = IC ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then A39: x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A31, TARSKI:def_1; A40: IC (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) = (IExec ((I ";" (Stop SCM+FSA)),P,s)) . (IC ) by A30 .= IC ((IExec (I,P,s)) +* (Start-At ((card I),SCM+FSA))) by A15, A16, SCMFSA8A:36 .= card I by FUNCT_4:113 ; thus (IExec ((if=0 (a,I,J)),P,s)) . x = (Result ((P +* (if=0 (a,I,J))),(Initialized s))) . x by A32 .= (Comput ((P +* (if=0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1))) . x by A20, A27, EXTPRO_1:23 .= IC (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))) by A38, EXTPRO_1:4 .= (IC (Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)))))) + ((card J) + 3) by A18, A10, A12, Th8, A14, A1 .= (IC (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)))) + ((card J) + 3) by A15, A16, EXTPRO_1:23, SCMFSA8A:34 .= (Start-At (((card I) + ((card J) + 3)),SCM+FSA)) . (IC ) by A40, FUNCOP_1:72 .= ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A38, A39, FUNCT_4:13 ; ::_thesis: verum end; end; end; dom (IExec ((if=0 (a,I,J)),P,s)) = the carrier of SCM+FSA by PARTFUN1:def_2 .= dom ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) by PARTFUN1:def_2 ; hence IExec ((if=0 (a,I,J)),P,s) = (IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A29, FUNCT_1:2 .= ((IExec (I,P,s)) +* (Start-At ((card I),SCM+FSA))) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A15, A16, SCMFSA8A:36 .= (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by FUNCT_4:114 ; ::_thesis: verum end; theorem Th15: :: SCMFSA8B:15 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a <> 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a <> 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a <> 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a <> 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) let a be read-write Int-Location; ::_thesis: ( s . a <> 0 & J is_closed_on s,P & J is_halting_on s,P implies ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) ) set I1 = I ";" (Stop SCM+FSA); set JI2 = ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); set s2 = Initialize s; set P2 = P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)); A1: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) c= P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by FUNCT_4:25; set s3 = Initialize s; set P3 = P +* (if=0 (a,I,J)); set s4 = Comput ((P +* (if=0 (a,I,J))),(Initialize s),1); set s5 = Comput ((P +* (if=0 (a,I,J))),(Initialize s),2); set i = a =0_goto ((card J) + 3); IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; then A2: IC (Initialize s) = IC (Start-At (0,SCM+FSA)) by FUNCT_4:13 .= 0 by FUNCOP_1:72 ; A3: if=0 (a,I,J) c= P +* (if=0 (a,I,J)) by FUNCT_4:25; A4: if=0 (a,I,J) = (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a =0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a =0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (a =0_goto ((card J) + 3)) ";" ((J ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (Macro (a =0_goto ((card J) + 3))) ";" (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by SCMFSA6A:25 ; then Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),(card (Macro (a =0_goto ((card J) + 3))))) c= if=0 (a,I,J) by Lm1; then A5: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= if=0 (a,I,J) by COMPOS_1:56; A6: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= P +* (if=0 (a,I,J)) by A5, A3, XBOOLE_1:1; A7: not a in dom (Start-At (0,SCM+FSA)) by SCMFSA_2:102; A8: 0 in dom (if=0 (a,I,J)) by Lm2; A9: if=0 (a,I,J) c= P +* (if=0 (a,I,J)) by FUNCT_4:25; A10: (P +* (if=0 (a,I,J))) . 0 = (if=0 (a,I,J)) . 0 by A8, FUNCT_4:13 .= a =0_goto ((card J) + 3) by Lm3 ; A11: 1 in dom (if=0 (a,I,J)) by Lm2; A12: Comput ((P +* (if=0 (a,I,J))),(Initialize s),(0 + 1)) = Following ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),0))) by EXTPRO_1:3 .= Following ((P +* (if=0 (a,I,J))),(Initialize s)) .= Exec ((a =0_goto ((card J) + 3)),(Initialize s)) by A2, A10, PBOOLE:143 ; assume s . a <> 0 ; ::_thesis: ( not J is_closed_on s,P or not J is_halting_on s,P or ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) ) then (Initialize s) . a <> 0 by A7, FUNCT_4:11; then A13: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) = succ 0 by A2, A12, SCMFSA_2:70 .= 0 + 1 ; A14: (P +* (if=0 (a,I,J))) . 1 = (if=0 (a,I,J)) . 1 by A11, A9, GRFUNC_1:2 .= goto 2 by Lm3 ; assume A15: J is_closed_on s,P ; ::_thesis: ( not J is_halting_on s,P or ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) ) A16: Comput ((P +* (if=0 (a,I,J))),(Initialize s),(1 + 1)) = Following ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),1))) by EXTPRO_1:3 .= Exec ((goto 2),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),1))) by A13, A14, PBOOLE:143 ; A17: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Initialize_s)_._f_=_(Comput_((P_+*_(if=0_(a,I,J))),(Initialize_s),2))_._f let f be FinSeq-Location ; ::_thesis: (Initialize s) . f = (Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)) . f thus (Initialize s) . f = (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) . f by A12, SCMFSA_2:70 .= (Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)) . f by A16, SCMFSA_2:69 ; ::_thesis: verum end; now__::_thesis:_for_a_being_Int-Location_holds_(Initialize_s)_._a_=_(Comput_((P_+*_(if=0_(a,I,J))),(Initialize_s),2))_._a let a be Int-Location; ::_thesis: (Initialize s) . a = (Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)) . a thus (Initialize s) . a = (Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)) . a by A12, SCMFSA_2:70 .= (Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)) . a by A16, SCMFSA_2:69 ; ::_thesis: verum end; then A18: DataPart (Initialize s) = DataPart (Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)) by A17, SCMFSA_M:2; assume A19: J is_halting_on s,P ; ::_thesis: ( if=0 (a,I,J) is_closed_on s,P & if=0 (a,I,J) is_halting_on s,P ) then A20: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on s,P by A15, SCMFSA8A:37; then A21: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on Initialize s,P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by Th7; A22: P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) halts_on Initialize s by A15, A19, SCMFSA8A:38; A23: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)) = 2 by A16, SCMFSA_2:69; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P_+*_(if=0_(a,I,J))),(Initialize_s),k))_in_dom_(if=0_(a,I,J)) let k be Element of NAT ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) ( k = 0 or 0 + 1 < k + 1 ) by XREAL_1:6; then A24: ( k = 0 or 1 <= k ) by NAT_1:13; percases ( 1 < k or k = 0 or k = 1 ) by A24, XXREAL_0:1; supposeA25: 1 < k ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) then consider k1 being Nat such that A26: k1 + 1 = k by NAT_1:6; reconsider k1 = k1 as Element of NAT by ORDINAL1:def_12; 0 + 1 < k1 + 1 by A25, A26; then consider k2 being Nat such that A27: k2 + 1 = k1 by NAT_1:6; reconsider k2 = k2 as Element of NAT by ORDINAL1:def_12; reconsider m = IC (Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s),k2)) as Element of NAT ; A28: card (if=0 (a,I,J)) = (card (Macro (a =0_goto ((card J) + 3)))) + (card (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))) by A4, SCMFSA6A:21 .= 2 + (card (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))) by COMPOS_1:56 ; m in dom (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by A20, SCMFSA7B:def_6; then m < card (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by AFINSQ_1:66; then A29: m + 2 < card (if=0 (a,I,J)) by A28, XREAL_1:6; IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) = IC (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)),k1)) by A26, EXTPRO_1:4 .= IC (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),1)),1)),k2)) by A27, EXTPRO_1:4 .= IC (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),(1 + 1))),k2)) by EXTPRO_1:4 .= m + 2 by A21, A23, A18, Th8, A6, A1 ; hence IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) in dom (if=0 (a,I,J)) by A29, AFINSQ_1:66; ::_thesis: verum end; suppose k = 0 ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) hence IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) in dom (if=0 (a,I,J)) by A8, A2, EXTPRO_1:2; ::_thesis: verum end; suppose k = 1 ; ::_thesis: IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),b1)) in dom (if=0 (a,I,J)) hence IC (Comput ((P +* (if=0 (a,I,J))),(Initialize s),k)) in dom (if=0 (a,I,J)) by A13, Lm2; ::_thesis: verum end; end; end; hence if=0 (a,I,J) is_closed_on s,P by SCMFSA7B:def_6; ::_thesis: if=0 (a,I,J) is_halting_on s,P CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s))) + 2)))) = CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialize s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s))))))),2) by A21, A23, A18, Th8, A6, A1 .= IncAddr ((halt SCM+FSA),2) by A22, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then P +* (if=0 (a,I,J)) halts_on Initialize s by EXTPRO_1:29; hence if=0 (a,I,J) is_halting_on s,P by SCMFSA7B:def_7; ::_thesis: verum end; theorem Th16: :: SCMFSA8B:16 for P being Instruction-Sequence of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location for s being State of SCM+FSA st s . a <> 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location for s being State of SCM+FSA st s . a <> 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location for s being State of SCM+FSA st s . a <> 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let a be read-write Int-Location; ::_thesis: for s being State of SCM+FSA st s . a <> 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: ( s . a <> 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) set I1 = I ";" (Stop SCM+FSA); set JI2 = ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); set s2 = Initialized s; set P2 = P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)); A1: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) c= P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by FUNCT_4:25; set P3 = P +* (if=0 (a,I,J)); set s4 = Comput ((P +* (if=0 (a,I,J))),(Initialized s),1); set s5 = Comput ((P +* (if=0 (a,I,J))),(Initialized s),2); set i = a =0_goto ((card J) + 3); 0 in dom (if=0 (a,I,J)) by Lm2; then A2: (P +* (if=0 (a,I,J))) . 0 = (if=0 (a,I,J)) . 0 by FUNCT_4:13 .= a =0_goto ((card J) + 3) by Lm3 ; A3: 1 in dom (if=0 (a,I,J)) by Lm2; A4: (P +* (if=0 (a,I,J))) . 1 = (if=0 (a,I,J)) . 1 by A3, FUNCT_4:13 .= goto 2 by Lm3 ; A5: dom (Initialize ((intloc 0) .--> 1)) = {(intloc 0),(IC )} by SCMFSA_M:11; ( a <> intloc 0 & a <> IC ) by SCMFSA_2:56; then A6: not a in dom (Initialize ((intloc 0) .--> 1)) by A5, TARSKI:def_2; IC in dom (Initialize ((intloc 0) .--> 1)) by MEMSTR_0:48; then A7: IC (Initialized s) = IC (Initialize ((intloc 0) .--> 1)) by FUNCT_4:13 .= 0 by MEMSTR_0:def_11 ; if=0 (a,I,J) = (((a =0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a =0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a =0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (a =0_goto ((card J) + 3)) ";" ((J ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (Macro (a =0_goto ((card J) + 3))) ";" (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by SCMFSA6A:25 ; then Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),(card (Macro (a =0_goto ((card J) + 3))))) c= if=0 (a,I,J) by Lm1; then A8: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= if=0 (a,I,J) by COMPOS_1:56; if=0 (a,I,J) c= P +* (if=0 (a,I,J)) by FUNCT_4:25; then A9: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= P +* (if=0 (a,I,J)) by A8, XBOOLE_1:1; A10: Comput ((P +* (if=0 (a,I,J))),(Initialized s),(0 + 1)) = Following ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),0))) by EXTPRO_1:3 .= Following ((P +* (if=0 (a,I,J))),(Initialized s)) .= Exec ((a =0_goto ((card J) + 3)),(Initialized s)) by A7, A2, PBOOLE:143 ; assume s . a <> 0 ; ::_thesis: ( not J is_closed_on Initialized s,P or not J is_halting_on Initialized s,P or IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) then (Initialized s) . a <> 0 by A6, FUNCT_4:11; then A11: IC (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) = succ 0 by A7, A10, SCMFSA_2:70 .= 0 + 1 ; assume A12: J is_closed_on Initialized s,P ; ::_thesis: ( not J is_halting_on Initialized s,P or IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) A13: Comput ((P +* (if=0 (a,I,J))),(Initialized s),(1 + 1)) = Following ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),1))) by EXTPRO_1:3 .= Exec ((goto 2),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),1))) by A11, A4, PBOOLE:143 ; then A14: IC (Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)) = 2 by SCMFSA_2:69; A15: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Initialized_s)_._f_=_(Comput_((P_+*_(if=0_(a,I,J))),(Initialized_s),2))_._f let f be FinSeq-Location ; ::_thesis: (Initialized s) . f = (Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)) . f thus (Initialized s) . f = (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) . f by A10, SCMFSA_2:70 .= (Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)) . f by A13, SCMFSA_2:69 ; ::_thesis: verum end; now__::_thesis:_for_a_being_Int-Location_holds_(Initialized_s)_._a_=_(Comput_((P_+*_(if=0_(a,I,J))),(Initialized_s),2))_._a let a be Int-Location; ::_thesis: (Initialized s) . a = (Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)) . a thus (Initialized s) . a = (Comput ((P +* (if=0 (a,I,J))),(Initialized s),1)) . a by A10, SCMFSA_2:70 .= (Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)) . a by A13, SCMFSA_2:69 ; ::_thesis: verum end; then A16: DataPart (Initialized s) = DataPart (Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)) by A15, SCMFSA_M:2; assume A17: J is_halting_on Initialized s,P ; ::_thesis: IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) then A18: P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) halts_on Initialized s by A12, SCMFSA8A:39; ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on Initialized s,P by A12, A17, SCMFSA8A:37; then A19: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on Initialized s,P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by Th6; A20: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2)))) = CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))))),2) by A19, A9, A14, A16, Th8, A1 .= IncAddr ((halt SCM+FSA),2) by A18, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then A21: P +* (if=0 (a,I,J)) halts_on Initialized s by EXTPRO_1:29; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_(LifeSpan_((P_+*_(((J_";"_(Goto_((card_I)_+_1)))_";"_I)_";"_(Stop_SCM+FSA))),(Initialized_s)))_+_2_holds_ CurInstr_((P_+*_(if=0_(a,I,J))),(Comput_((P_+*_(if=0_(a,I,J))),(Initialized_s),l)))_<>_halt_SCM+FSA let l be Element of NAT ; ::_thesis: ( l < (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 implies CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA ) assume A22: l < (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 ; ::_thesis: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA A23: Comput ((P +* (if=0 (a,I,J))),(Initialized s),0) = Initialized s ; percases ( l = 0 or l = 1 or ( l <> 0 & l <> 1 ) ) ; suppose l = 0 ; ::_thesis: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA hence CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) <> halt SCM+FSA by A7, A2, A23, PBOOLE:143; ::_thesis: verum end; suppose l = 1 ; ::_thesis: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA hence CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) <> halt SCM+FSA by A11, A4, PBOOLE:143; ::_thesis: verum end; supposeA24: ( l <> 0 & l <> 1 ) ; ::_thesis: not CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),b1))) = halt SCM+FSA assume A25: CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA ; ::_thesis: contradiction consider n being Nat such that A26: l = n + 1 by A24, NAT_1:6; n <> 0 by A24, A26; then consider l2 being Nat such that A27: n = l2 + 1 by NAT_1:6; reconsider l2 = l2 as Element of NAT by ORDINAL1:def_12; A28: InsCode (halt SCM+FSA) = 0 by COMPOS_1:70; A29: Comput ((P +* (if=0 (a,I,J))),(Initialized s),(l2 + (1 + 1))) = Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),(1 + 1))),l2) by EXTPRO_1:4; InsCode (CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),l2)))) = InsCode (IncAddr ((CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),l2)))),2)) by COMPOS_0:def_9 .= 0 by A26, A27, A25, A29, A19, A14, A16, Th8, A9, A1, A28 ; then A30: CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),l2))) = halt SCM+FSA by SCMFSA_2:95; n + 1 < ((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 1) + 1 by A22, A26; then n < (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 1 by XREAL_1:6; then l2 < LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)) by A27, XREAL_1:6; hence contradiction by A18, A30, EXTPRO_1:def_15; ::_thesis: verum end; end; end; then for l being Element of NAT st CurInstr ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA holds (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 <= l ; then A31: LifeSpan ((P +* (if=0 (a,I,J))),(Initialized s)) = (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 by A20, A21, EXTPRO_1:def_15; A32: DataPart (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) = DataPart (Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))) by A17, A12, EXTPRO_1:23, SCMFSA8A:39 .= DataPart (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))) by A19, A14, A16, Th8, A9, A1 .= DataPart (Comput ((P +* (if=0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2))) by EXTPRO_1:4 .= DataPart (Result ((P +* (if=0 (a,I,J))),(Initialized s))) by A21, A31, EXTPRO_1:23 ; A33: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((if=0_(a,I,J)),P,s))_holds_ (IExec_((if=0_(a,I,J)),P,s))_._x_=_((IExec_((((J_";"_(Goto_((card_I)_+_1)))_";"_I)_";"_(Stop_SCM+FSA)),P,s))_+*_(Start-At_((((card_I)_+_(card_J))_+_3),SCM+FSA)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((if=0 (a,I,J)),P,s)) implies (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 ) A34: dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) = {(IC )} by FUNCOP_1:13; A35: IExec ((if=0 (a,I,J)),P,s) = Result ((P +* (if=0 (a,I,J))),(Initialized s)) by SCMFSA6B:def_1; A36: IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s) = Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)) by SCMFSA6B:def_1; assume A37: x in dom (IExec ((if=0 (a,I,J)),P,s)) ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 percases ( x is Int-Location or x is FinSeq-Location or x = IC ) by A37, SCMFSA_M:1; supposeA38: x is Int-Location ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:56; then A39: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A34, TARSKI:def_1; thus (IExec ((if=0 (a,I,J)),P,s)) . x = (Result ((P +* (if=0 (a,I,J))),(Initialized s))) . x by A35 .= (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) . x by A32, A38, SCMFSA_M:2 .= (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) . x by A36 .= ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A39, FUNCT_4:11 ; ::_thesis: verum end; supposeA40: x is FinSeq-Location ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:57; then A41: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A34, TARSKI:def_1; thus (IExec ((if=0 (a,I,J)),P,s)) . x = (Result ((P +* (if=0 (a,I,J))),(Initialized s))) . x by A35 .= (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) . x by A32, A40, SCMFSA_M:2 .= (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) . x by A36 .= ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A41, FUNCT_4:11 ; ::_thesis: verum end; supposeA42: x = IC ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then A43: x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A34, TARSKI:def_1; A44: IC (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) = IC (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) by A36 .= ((card I) + (card J)) + 1 by A12, A17, SCMFSA8A:40 ; thus (IExec ((if=0 (a,I,J)),P,s)) . x = (Result ((P +* (if=0 (a,I,J))),(Initialized s))) . x by A35 .= (Comput ((P +* (if=0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2))) . x by A21, A31, EXTPRO_1:23 .= IC (Comput ((P +* (if=0 (a,I,J))),(Comput ((P +* (if=0 (a,I,J))),(Initialized s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))) by A42, EXTPRO_1:4 .= (IC (Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)))))) + 2 by A19, A14, A16, Th8, A9, A1 .= (IC (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)))) + 2 by A17, A12, EXTPRO_1:23, SCMFSA8A:39 .= (Start-At (((((card I) + (card J)) + 1) + 2),SCM+FSA)) . (IC ) by A44, FUNCOP_1:72 .= ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A42, A43, FUNCT_4:13 ; ::_thesis: verum end; end; end; dom (IExec ((if=0 (a,I,J)),P,s)) = the carrier of SCM+FSA by PARTFUN1:def_2 .= dom ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) by PARTFUN1:def_2 ; hence IExec ((if=0 (a,I,J)),P,s) = (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A33, FUNCT_1:2 .= ((IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 1),SCM+FSA))) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A12, A17, SCMFSA8A:41 .= (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by FUNCT_4:114 ; ::_thesis: verum end; theorem Th17: :: SCMFSA8B:17 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if=0 (a,I,J) is parahalting & ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if=0 (a,I,J) is parahalting & ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if=0 (a,I,J) is parahalting & ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a being read-write Int-Location holds ( if=0 (a,I,J) is parahalting & ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) let a be read-write Int-Location; ::_thesis: ( if=0 (a,I,J) is parahalting & ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) A1: I is_halting_on Initialized s,P by SCMFSA7B:19; for s being 0 -started State of SCM+FSA for P being Instruction-Sequence of SCM+FSA st if=0 (a,I,J) c= P holds P halts_on s proof let s be 0 -started State of SCM+FSA; ::_thesis: for P being Instruction-Sequence of SCM+FSA st if=0 (a,I,J) c= P holds P halts_on s Start-At (0,SCM+FSA) c= s by MEMSTR_0:29; then A2: s = Initialize s by FUNCT_4:98; let Q be Instruction-Sequence of SCM+FSA; ::_thesis: ( if=0 (a,I,J) c= Q implies Q halts_on s ) assume A3: if=0 (a,I,J) c= Q ; ::_thesis: Q halts_on s A4: if=0 (a,I,J) c= Q by A3; A5: I is_closed_on s,Q by SCMFSA7B:18; A6: I is_halting_on s,Q by SCMFSA7B:19; A7: J is_halting_on s,Q by SCMFSA7B:19; A8: J is_closed_on s,Q by SCMFSA7B:18; A9: Q +* (if=0 (a,I,J)) = Q by A4, FUNCT_4:98; percases ( s . a = 0 or s . a <> 0 ) ; suppose s . a = 0 ; ::_thesis: Q halts_on s then if=0 (a,I,J) is_halting_on s,Q by A5, A6, Th13; hence Q halts_on s by A2, A9, SCMFSA7B:def_7; ::_thesis: verum end; suppose s . a <> 0 ; ::_thesis: Q halts_on s then if=0 (a,I,J) is_halting_on s,Q by A8, A7, Th15; hence Q halts_on s by A2, A9, SCMFSA7B:def_7; ::_thesis: verum end; end; end; hence if=0 (a,I,J) is parahalting by AMISTD_1:def_11; ::_thesis: ( ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) A10: J is_closed_on Initialized s,P by SCMFSA7B:18; I is_closed_on Initialized s,P by SCMFSA7B:18; hence ( s . a = 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) by A1, Th14; ::_thesis: ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) J is_halting_on Initialized s,P by SCMFSA7B:19; hence ( s . a <> 0 implies IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) by Th16, A10; ::_thesis: verum end; theorem Th18: :: SCMFSA8B:18 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a = 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a = 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a = 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a being read-write Int-Location holds ( IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a = 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let a be read-write Int-Location; ::_thesis: ( IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a = 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) hereby ::_thesis: ( ( s . a = 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) percases ( s . a = 0 or s . a <> 0 ) ; suppose s . a = 0 ; ::_thesis: IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 then IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th17; hence IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 by FUNCT_4:113; ::_thesis: verum end; suppose s . a <> 0 ; ::_thesis: IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 then IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th17; hence IC (IExec ((if=0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 by FUNCT_4:113; ::_thesis: verum end; end; end; hereby ::_thesis: ( s . a <> 0 implies ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) assume s . a = 0 ; ::_thesis: ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) then A1: IExec ((if=0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th17; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f let d be Int-Location; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d not d in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:102; hence (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d by A1, FUNCT_4:11; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f not f in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:103; hence (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f by A1, FUNCT_4:11; ::_thesis: verum end; assume s . a <> 0 ; ::_thesis: ( ( for d being Int-Location holds (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) then A2: IExec ((if=0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th17; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f let d be Int-Location; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d not d in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:102; hence (IExec ((if=0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d by A2, FUNCT_4:11; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f not f in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:103; hence (IExec ((if=0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f by A2, FUNCT_4:11; ::_thesis: verum end; theorem Th19: :: SCMFSA8B:19 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a > 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) let a be read-write Int-Location; ::_thesis: ( s . a > 0 & I is_closed_on s,P & I is_halting_on s,P implies ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) ) set I1 = I ";" (Stop SCM+FSA); set s1 = Initialize s; set P1 = P +* (I ";" (Stop SCM+FSA)); set s3 = Initialize s; set P3 = P +* (if>0 (a,I,J)); set s4 = Comput ((P +* (if>0 (a,I,J))),(Initialize s),1); set i = a >0_goto ((card J) + 3); A1: I ";" (Stop SCM+FSA) c= P +* (I ";" (Stop SCM+FSA)) by FUNCT_4:25; A2: not a in dom (Start-At (0,SCM+FSA)) by SCMFSA_2:102; A3: 0 in dom (if>0 (a,I,J)) by Lm2; A4: (P +* (if>0 (a,I,J))) . 0 = (if>0 (a,I,J)) . 0 by A3, FUNCT_4:13 .= a >0_goto ((card J) + 3) by Lm3 ; IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; then A5: IC (Initialize s) = IC (Start-At (0,SCM+FSA)) by FUNCT_4:13 .= 0 by FUNCOP_1:72 ; A6: if>0 (a,I,J) c= P +* (if>0 (a,I,J)) by FUNCT_4:25; A7: if>0 (a,I,J) = (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25; card (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) = (card ((Macro (a >0_goto ((card J) + 3))) ";" J)) + (card (Goto ((card I) + 1))) by SCMFSA6A:21 .= (card ((Macro (a >0_goto ((card J) + 3))) ";" J)) + 1 by SCMFSA8A:15 .= ((card (Macro (a >0_goto ((card J) + 3)))) + (card J)) + 1 by SCMFSA6A:21 .= ((card J) + 2) + 1 by COMPOS_1:56 .= (card J) + (2 + 1) ; then A8: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= if>0 (a,I,J) by A7, Lm1; A9: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= P +* (if>0 (a,I,J)) by A8, A6, XBOOLE_1:1; A10: Comput ((P +* (if>0 (a,I,J))),(Initialize s),(0 + 1)) = Following ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),0))) by EXTPRO_1:3 .= Following ((P +* (if>0 (a,I,J))),(Initialize s)) .= Exec ((a >0_goto ((card J) + 3)),(Initialize s)) by A5, A4, PBOOLE:143 ; A11: for f being FinSeq-Location holds (Initialize s) . f = (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) . f by A10, SCMFSA_2:71; for a being Int-Location holds (Initialize s) . a = (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) . a by A10, SCMFSA_2:71; then A12: DataPart (Initialize s) = DataPart (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) by A11, SCMFSA_M:2; assume s . a > 0 ; ::_thesis: ( not I is_closed_on s,P or not I is_halting_on s,P or ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) ) then (Initialize s) . a > 0 by A2, FUNCT_4:11; then A13: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) = (card J) + 3 by A10, SCMFSA_2:71; assume A14: I is_closed_on s,P ; ::_thesis: ( not I is_halting_on s,P or ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) ) assume A15: I is_halting_on s,P ; ::_thesis: ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) then A16: I ";" (Stop SCM+FSA) is_closed_on s,P by A14, SCMFSA8A:30; I ";" (Stop SCM+FSA) is_halting_on s,P by A14, A15, SCMFSA8A:30; then A17: P +* (I ";" (Stop SCM+FSA)) halts_on Initialize s by SCMFSA7B:def_7; DataPart s = DataPart (Initialize s) by MEMSTR_0:79; then A18: I ";" (Stop SCM+FSA) is_closed_on Initialize s,P +* (I ";" (Stop SCM+FSA)) by A16, Th3; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P_+*_(if>0_(a,I,J))),(Initialize_s),k))_in_dom_(if>0_(a,I,J)) let k be Element of NAT ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) percases ( 0 < k or k = 0 ) ; suppose 0 < k ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) then consider k1 being Nat such that A19: k1 + 1 = k by NAT_1:6; reconsider k1 = k1 as Element of NAT by ORDINAL1:def_12; reconsider m = IC (Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialize s),k1)) as Element of NAT ; m in dom (I ";" (Stop SCM+FSA)) by A16, SCMFSA7B:def_6; then A20: m < card (I ";" (Stop SCM+FSA)) by AFINSQ_1:66; card (Stop SCM+FSA) = 1 by COMPOS_1:4; then A21: card (I ";" (Stop SCM+FSA)) = (card I) + 1 by SCMFSA6A:21; card (if>0 (a,I,J)) = ((card I) + (card J)) + 4 by Th12 .= ((card J) + 3) + (card (I ";" (Stop SCM+FSA))) by A21 ; then A22: m + ((card J) + 3) < card (if>0 (a,I,J)) by A20, XREAL_1:6; IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) = IC (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)),k1)) by A19, EXTPRO_1:4 .= m + ((card J) + 3) by A18, A13, A12, Th8, A9, A1 ; hence IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) in dom (if>0 (a,I,J)) by A22, AFINSQ_1:66; ::_thesis: verum end; suppose k = 0 ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) hence IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) in dom (if>0 (a,I,J)) by A3, A5, EXTPRO_1:2; ::_thesis: verum end; end; end; hence if>0 (a,I,J) is_closed_on s,P by SCMFSA7B:def_6; ::_thesis: if>0 (a,I,J) is_halting_on s,P CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialize s))) + 1)))) = CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialize s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialize s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialize s))))))),((card J) + 3)) by A18, A13, A12, Th8, A9, A1 .= IncAddr ((halt SCM+FSA),((card J) + 3)) by A17, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then P +* (if>0 (a,I,J)) halts_on Initialize s by EXTPRO_1:29; hence if>0 (a,I,J) is_halting_on s,P by SCMFSA7B:def_7; ::_thesis: verum end; theorem Th20: :: SCMFSA8B:20 for P being Instruction-Sequence of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location for s being State of SCM+FSA st s . a > 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location for s being State of SCM+FSA st s . a > 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location for s being State of SCM+FSA st s . a > 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let a be read-write Int-Location; ::_thesis: for s being State of SCM+FSA st s . a > 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: ( s . a > 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) set I1 = I ";" (Stop SCM+FSA); set s1 = Initialized s; set P1 = P +* (I ";" (Stop SCM+FSA)); set P3 = P +* (if>0 (a,I,J)); set s4 = Comput ((P +* (if>0 (a,I,J))),(Initialized s),1); set i = a >0_goto ((card J) + 3); A1: I ";" (Stop SCM+FSA) c= P +* (I ";" (Stop SCM+FSA)) by FUNCT_4:25; A2: if>0 (a,I,J) = (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25; A3: 0 in dom (if>0 (a,I,J)) by Lm2; A4: (P +* (if>0 (a,I,J))) . 0 = (if>0 (a,I,J)) . 0 by A3, FUNCT_4:13 .= a >0_goto ((card J) + 3) by Lm3 ; A5: dom (Initialize ((intloc 0) .--> 1)) = {(intloc 0),(IC )} by SCMFSA_M:11; ( a <> intloc 0 & a <> IC ) by SCMFSA_2:56; then A6: not a in dom (Initialize ((intloc 0) .--> 1)) by A5, TARSKI:def_2; IC in dom (Initialize ((intloc 0) .--> 1)) by MEMSTR_0:48; then A7: IC (Initialized s) = IC (Initialize ((intloc 0) .--> 1)) by FUNCT_4:13 .= 0 by MEMSTR_0:def_11 ; A8: Comput ((P +* (if>0 (a,I,J))),(Initialized s),(0 + 1)) = Following ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),0))) by EXTPRO_1:3 .= Following ((P +* (if>0 (a,I,J))),(Initialized s)) .= Exec ((a >0_goto ((card J) + 3)),(Initialized s)) by A7, A4, PBOOLE:143 ; A9: if>0 (a,I,J) c= P +* (if>0 (a,I,J)) by FUNCT_4:25; assume s . a > 0 ; ::_thesis: ( not I is_closed_on Initialized s,P or not I is_halting_on Initialized s,P or IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) then (Initialized s) . a > 0 by A6, FUNCT_4:11; then A10: IC (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) = (card J) + 3 by A8, SCMFSA_2:71; A11: for f being FinSeq-Location holds (Initialized s) . f = (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) . f by A8, SCMFSA_2:71; for a being Int-Location holds (Initialized s) . a = (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) . a by A8, SCMFSA_2:71; then A12: DataPart (Initialized s) = DataPart (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) by A11, SCMFSA_M:2; card (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) = (card ((Macro (a >0_goto ((card J) + 3))) ";" J)) + (card (Goto ((card I) + 1))) by SCMFSA6A:21 .= (card ((Macro (a >0_goto ((card J) + 3))) ";" J)) + 1 by SCMFSA8A:15 .= ((card (Macro (a >0_goto ((card J) + 3)))) + (card J)) + 1 by SCMFSA6A:21 .= ((card J) + 2) + 1 by COMPOS_1:56 .= (card J) + (2 + 1) ; then A13: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= if>0 (a,I,J) by A2, Lm1; A14: Reloc ((I ";" (Stop SCM+FSA)),((card J) + 3)) c= P +* (if>0 (a,I,J)) by A13, A9, XBOOLE_1:1; assume A15: I is_closed_on Initialized s,P ; ::_thesis: ( not I is_halting_on Initialized s,P or IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) assume A16: I is_halting_on Initialized s,P ; ::_thesis: IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) then A17: P +* (I ";" (Stop SCM+FSA)) halts_on Initialized s by A15, SCMFSA8A:34; I ";" (Stop SCM+FSA) is_closed_on Initialized s,P by A15, A16, SCMFSA8A:30; then A18: I ";" (Stop SCM+FSA) is_closed_on Initialized s,P +* (I ";" (Stop SCM+FSA)) by Th6; A19: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1)))) = CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))))),((card J) + 3)) by A18, A10, A12, Th8, A14, A1 .= IncAddr ((halt SCM+FSA),((card J) + 3)) by A17, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then A20: P +* (if>0 (a,I,J)) halts_on Initialized s by EXTPRO_1:29; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_(LifeSpan_((P_+*_(I_";"_(Stop_SCM+FSA))),(Initialized_s)))_+_1_holds_ CurInstr_((P_+*_(if>0_(a,I,J))),(Comput_((P_+*_(if>0_(a,I,J))),(Initialized_s),l)))_<>_halt_SCM+FSA let l be Element of NAT ; ::_thesis: ( l < (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 implies CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA ) assume A21: l < (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 ; ::_thesis: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA A22: Comput ((P +* (if>0 (a,I,J))),(Initialized s),0) = Initialized s ; percases ( l = 0 or l <> 0 ) ; suppose l = 0 ; ::_thesis: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA hence CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) <> halt SCM+FSA by A7, A4, A22, PBOOLE:143; ::_thesis: verum end; suppose l <> 0 ; ::_thesis: not CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) = halt SCM+FSA then consider n being Nat such that A23: l = n + 1 by NAT_1:6; assume A24: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA ; ::_thesis: contradiction reconsider n = n as Element of NAT by ORDINAL1:def_12; A25: InsCode (halt SCM+FSA) = 0 by COMPOS_1:70; InsCode (CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),n)))) = InsCode (IncAddr ((CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),n)))),((card J) + 3))) by COMPOS_0:def_9 .= InsCode (CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)),n)))) by A18, A10, A12, Th8, A14, A1 .= 0 by A23, A24, EXTPRO_1:4, A25 ; then A26: CurInstr ((P +* (I ";" (Stop SCM+FSA))),(Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),n))) = halt SCM+FSA by SCMFSA_2:95; n < LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)) by A21, A23, XREAL_1:6; hence contradiction by A17, A26, EXTPRO_1:def_15; ::_thesis: verum end; end; end; then for l being Element of NAT st CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA holds (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 <= l ; then A27: LifeSpan ((P +* (if>0 (a,I,J))),(Initialized s)) = (LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1 by A19, A20, EXTPRO_1:def_15; A28: DataPart (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) = DataPart (Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))) by A15, A16, EXTPRO_1:23, SCMFSA8A:34 .= DataPart (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))) by A18, A10, A12, Th8, A14, A1 .= DataPart (Comput ((P +* (if>0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1))) by EXTPRO_1:4 .= DataPart (Result ((P +* (if>0 (a,I,J))),(Initialized s))) by A20, A27, EXTPRO_1:23 ; A29: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((if>0_(a,I,J)),P,s))_holds_ (IExec_((if>0_(a,I,J)),P,s))_._x_=_((IExec_((I_";"_(Stop_SCM+FSA)),P,s))_+*_(Start-At_((((card_I)_+_(card_J))_+_3),SCM+FSA)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((if>0 (a,I,J)),P,s)) implies (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 ) A30: dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) = {(IC )} by FUNCOP_1:13; A31: IExec ((if>0 (a,I,J)),P,s) = Result ((P +* (if>0 (a,I,J))),(Initialized s)) by SCMFSA6B:def_1; A32: IExec ((I ";" (Stop SCM+FSA)),P,s) = Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)) by SCMFSA6B:def_1; assume A33: x in dom (IExec ((if>0 (a,I,J)),P,s)) ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 percases ( x is Int-Location or x is FinSeq-Location or x = IC ) by A33, SCMFSA_M:1; supposeA34: x is Int-Location ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:56; then A35: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A30, TARSKI:def_1; thus (IExec ((if>0 (a,I,J)),P,s)) . x = (Result ((P +* (if>0 (a,I,J))),(Initialized s))) . x by A31 .= (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) . x by A28, A34, SCMFSA_M:2 .= (IExec ((I ";" (Stop SCM+FSA)),P,s)) . x by A32 .= ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A35, FUNCT_4:11 ; ::_thesis: verum end; supposeA36: x is FinSeq-Location ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:57; then A37: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A30, TARSKI:def_1; thus (IExec ((if>0 (a,I,J)),P,s)) . x = (Result ((P +* (if>0 (a,I,J))),(Initialized s))) . x by A31 .= (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) . x by A28, A36, SCMFSA_M:2 .= (IExec ((I ";" (Stop SCM+FSA)),P,s)) . x by A32 .= ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A37, FUNCT_4:11 ; ::_thesis: verum end; supposeA38: x = IC ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then A39: x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A30, TARSKI:def_1; A40: IC (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) = (IExec ((I ";" (Stop SCM+FSA)),P,s)) . (IC ) by A32 .= IC ((IExec (I,P,s)) +* (Start-At ((card I),SCM+FSA))) by A15, A16, SCMFSA8A:36 .= card I by FUNCT_4:113 ; thus (IExec ((if>0 (a,I,J)),P,s)) . x = (Result ((P +* (if>0 (a,I,J))),(Initialized s))) . x by A31 .= (Comput ((P +* (if>0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))) + 1))) . x by A20, A27, EXTPRO_1:23 .= IC (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s))))) by A38, EXTPRO_1:4 .= (IC (Comput ((P +* (I ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)))))) + ((card J) + 3) by A18, A10, A12, Th8, A1, A14 .= (IC (Result ((P +* (I ";" (Stop SCM+FSA))),(Initialized s)))) + ((card J) + 3) by A15, A16, EXTPRO_1:23, SCMFSA8A:34 .= (Start-At (((card I) + ((card J) + 3)),SCM+FSA)) . (IC ) by A40, FUNCOP_1:72 .= ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A38, A39, FUNCT_4:13 ; ::_thesis: verum end; end; end; dom (IExec ((if>0 (a,I,J)),P,s)) = the carrier of SCM+FSA by PARTFUN1:def_2 .= dom ((IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) by PARTFUN1:def_2 ; hence IExec ((if>0 (a,I,J)),P,s) = (IExec ((I ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A29, FUNCT_1:2 .= ((IExec (I,P,s)) +* (Start-At ((card I),SCM+FSA))) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A15, A16, SCMFSA8A:36 .= (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by FUNCT_4:114 ; ::_thesis: verum end; theorem Th21: :: SCMFSA8B:21 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a <= 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a <= 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a <= 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a <= 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) let a be read-write Int-Location; ::_thesis: ( s . a <= 0 & J is_closed_on s,P & J is_halting_on s,P implies ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) ) set I1 = I ";" (Stop SCM+FSA); set JI2 = ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); set s2 = Initialize s; set P2 = P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)); A1: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) c= P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by FUNCT_4:25; set s3 = Initialize s; set P3 = P +* (if>0 (a,I,J)); set s4 = Comput ((P +* (if>0 (a,I,J))),(Initialize s),1); set s5 = Comput ((P +* (if>0 (a,I,J))),(Initialize s),2); set i = a >0_goto ((card J) + 3); IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; then A2: IC (Initialize s) = IC (Start-At (0,SCM+FSA)) by FUNCT_4:13 .= 0 by FUNCOP_1:72 ; A3: if>0 (a,I,J) c= P +* (if>0 (a,I,J)) by FUNCT_4:25; A4: if>0 (a,I,J) = (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a >0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a >0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (a >0_goto ((card J) + 3)) ";" ((J ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (Macro (a >0_goto ((card J) + 3))) ";" (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by SCMFSA6A:25 ; then A5: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),(card (Macro (a >0_goto ((card J) + 3))))) c= if>0 (a,I,J) by Lm1; card (Macro (a >0_goto ((card J) + 3))) = 2 by COMPOS_1:56; then A6: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= P +* (if>0 (a,I,J)) by A3, A5, XBOOLE_1:1; A7: not a in dom (Start-At (0,SCM+FSA)) by SCMFSA_2:102; A8: 0 in dom (if>0 (a,I,J)) by Lm2; A9: (P +* (if>0 (a,I,J))) . 0 = (if>0 (a,I,J)) . 0 by A8, FUNCT_4:13 .= a >0_goto ((card J) + 3) by Lm3 ; A10: 1 in dom (if>0 (a,I,J)) by Lm2; A11: Comput ((P +* (if>0 (a,I,J))),(Initialize s),(0 + 1)) = Following ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),0))) by EXTPRO_1:3 .= Following ((P +* (if>0 (a,I,J))),(Initialize s)) .= Exec ((a >0_goto ((card J) + 3)),(Initialize s)) by A2, A9, PBOOLE:143 ; assume s . a <= 0 ; ::_thesis: ( not J is_closed_on s,P or not J is_halting_on s,P or ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) ) then (Initialize s) . a <= 0 by A7, FUNCT_4:11; then A12: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) = succ 0 by A2, A11, SCMFSA_2:71 .= 0 + 1 ; A13: (P +* (if>0 (a,I,J))) . 1 = (if>0 (a,I,J)) . 1 by A10, FUNCT_4:13 .= goto 2 by Lm3 ; assume A14: J is_closed_on s,P ; ::_thesis: ( not J is_halting_on s,P or ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) ) A15: Comput ((P +* (if>0 (a,I,J))),(Initialize s),(1 + 1)) = Following ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),1))) by EXTPRO_1:3 .= Exec ((goto 2),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),1))) by A12, A13, PBOOLE:143 ; A16: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Initialize_s)_._f_=_(Comput_((P_+*_(if>0_(a,I,J))),(Initialize_s),2))_._f let f be FinSeq-Location ; ::_thesis: (Initialize s) . f = (Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)) . f thus (Initialize s) . f = (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) . f by A11, SCMFSA_2:71 .= (Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)) . f by A15, SCMFSA_2:69 ; ::_thesis: verum end; now__::_thesis:_for_a_being_Int-Location_holds_(Initialize_s)_._a_=_(Comput_((P_+*_(if>0_(a,I,J))),(Initialize_s),2))_._a let a be Int-Location; ::_thesis: (Initialize s) . a = (Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)) . a thus (Initialize s) . a = (Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)) . a by A11, SCMFSA_2:71 .= (Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)) . a by A15, SCMFSA_2:69 ; ::_thesis: verum end; then A17: DataPart (Initialize s) = DataPart (Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)) by A16, SCMFSA_M:2; assume A18: J is_halting_on s,P ; ::_thesis: ( if>0 (a,I,J) is_closed_on s,P & if>0 (a,I,J) is_halting_on s,P ) then A19: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on s,P by A14, SCMFSA8A:37; then A20: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on Initialize s,P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by Th7; A21: P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) halts_on Initialize s by A14, A18, SCMFSA8A:38; A22: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)) = 2 by A15, SCMFSA_2:69; now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P_+*_(if>0_(a,I,J))),(Initialize_s),k))_in_dom_(if>0_(a,I,J)) let k be Element of NAT ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) ( k = 0 or 0 + 1 < k + 1 ) by XREAL_1:6; then A23: ( k = 0 or 1 <= k ) by NAT_1:13; percases ( 1 < k or k = 0 or k = 1 ) by A23, XXREAL_0:1; supposeA24: 1 < k ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) then consider k1 being Nat such that A25: k1 + 1 = k by NAT_1:6; 0 + 1 < k1 + 1 by A24, A25; then consider k2 being Nat such that A26: k2 + 1 = k1 by NAT_1:6; reconsider k1 = k1, k2 = k2 as Element of NAT by ORDINAL1:def_12; reconsider m = IC (Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s),k2)) as Element of NAT ; A27: card (if>0 (a,I,J)) = (card (Macro (a >0_goto ((card J) + 3)))) + (card (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))) by A4, SCMFSA6A:21 .= 2 + (card (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))) by COMPOS_1:56 ; m in dom (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by A19, SCMFSA7B:def_6; then m < card (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by AFINSQ_1:66; then A28: m + 2 < card (if>0 (a,I,J)) by A27, XREAL_1:6; IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) = IC (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)),k1)) by A25, EXTPRO_1:4 .= IC (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),1)),1)),k2)) by A26, EXTPRO_1:4 .= IC (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),(1 + 1))),k2)) by EXTPRO_1:4 .= m + 2 by A20, A22, A17, Th8, A6, A1 ; hence IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) in dom (if>0 (a,I,J)) by A28, AFINSQ_1:66; ::_thesis: verum end; suppose k = 0 ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) hence IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) in dom (if>0 (a,I,J)) by A8, A2, EXTPRO_1:2; ::_thesis: verum end; suppose k = 1 ; ::_thesis: IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),b1)) in dom (if>0 (a,I,J)) hence IC (Comput ((P +* (if>0 (a,I,J))),(Initialize s),k)) in dom (if>0 (a,I,J)) by A12, Lm2; ::_thesis: verum end; end; end; hence if>0 (a,I,J) is_closed_on s,P by SCMFSA7B:def_6; ::_thesis: if>0 (a,I,J) is_halting_on s,P CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s))) + 2)))) = CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialize s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialize s))))))),2) by A20, A6, A22, A17, Th8, A1 .= IncAddr ((halt SCM+FSA),2) by A21, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then P +* (if>0 (a,I,J)) halts_on Initialize s by EXTPRO_1:29; hence if>0 (a,I,J) is_halting_on s,P by SCMFSA7B:def_7; ::_thesis: verum end; theorem Th22: :: SCMFSA8B:22 for P being Instruction-Sequence of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location for s being State of SCM+FSA st s . a <= 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location for s being State of SCM+FSA st s . a <= 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location for s being State of SCM+FSA st s . a <= 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let a be read-write Int-Location; ::_thesis: for s being State of SCM+FSA st s . a <= 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: ( s . a <= 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) set I1 = I ";" (Stop SCM+FSA); set JI2 = ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA); set s2 = Initialized s; set P2 = P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)); A1: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) c= P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by FUNCT_4:25; set P3 = P +* (if>0 (a,I,J)); set s4 = Comput ((P +* (if>0 (a,I,J))),(Initialized s),1); set s5 = Comput ((P +* (if>0 (a,I,J))),(Initialized s),2); set i = a >0_goto ((card J) + 3); 0 in dom (if>0 (a,I,J)) by Lm2; then A2: (P +* (if>0 (a,I,J))) . 0 = (if>0 (a,I,J)) . 0 by FUNCT_4:13 .= a >0_goto ((card J) + 3) by Lm3 ; A3: 1 in dom (if>0 (a,I,J)) by Lm2; A4: (P +* (if>0 (a,I,J))) . 1 = (if>0 (a,I,J)) . 1 by A3, FUNCT_4:13 .= goto 2 by Lm3 ; A5: dom (Initialize ((intloc 0) .--> 1)) = {(intloc 0),(IC )} by SCMFSA_M:11; ( a <> intloc 0 & a <> IC ) by SCMFSA_2:56; then A6: not a in dom (Initialize ((intloc 0) .--> 1)) by A5, TARSKI:def_2; IC in dom (Initialize ((intloc 0) .--> 1)) by MEMSTR_0:48; then A7: IC (Initialized s) = IC (Initialize ((intloc 0) .--> 1)) by FUNCT_4:13 .= 0 by MEMSTR_0:def_11 ; if>0 (a,I,J) = (((a >0_goto ((card J) + 3)) ";" J) ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA)) by SCMFSA6A:25 .= ((a >0_goto ((card J) + 3)) ";" J) ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (a >0_goto ((card J) + 3)) ";" (J ";" ((Goto ((card I) + 1)) ";" (I ";" (Stop SCM+FSA)))) by SCMFSA6A:29 .= (a >0_goto ((card J) + 3)) ";" ((J ";" (Goto ((card I) + 1))) ";" (I ";" (Stop SCM+FSA))) by SCMFSA6A:25 .= (Macro (a >0_goto ((card J) + 3))) ";" (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by SCMFSA6A:25 ; then Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),(card (Macro (a >0_goto ((card J) + 3))))) c= if>0 (a,I,J) by Lm1; then A8: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= if>0 (a,I,J) by COMPOS_1:56; if>0 (a,I,J) c= P +* (if>0 (a,I,J)) by FUNCT_4:25; then A9: Reloc ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),2) c= P +* (if>0 (a,I,J)) by A8, XBOOLE_1:1; A10: Comput ((P +* (if>0 (a,I,J))),(Initialized s),(0 + 1)) = Following ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),0))) by EXTPRO_1:3 .= Following ((P +* (if>0 (a,I,J))),(Initialized s)) .= Exec ((a >0_goto ((card J) + 3)),(Initialized s)) by A7, A2, PBOOLE:143 ; assume s . a <= 0 ; ::_thesis: ( not J is_closed_on Initialized s,P or not J is_halting_on Initialized s,P or IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) then (Initialized s) . a <= 0 by A6, FUNCT_4:11; then A11: IC (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) = succ 0 by A7, A10, SCMFSA_2:71 .= 0 + 1 ; assume A12: J is_closed_on Initialized s,P ; ::_thesis: ( not J is_halting_on Initialized s,P or IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) A13: Comput ((P +* (if>0 (a,I,J))),(Initialized s),(1 + 1)) = Following ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),1))) by EXTPRO_1:3 .= Exec ((goto 2),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),1))) by A11, A4, PBOOLE:143 ; then A14: IC (Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)) = 2 by SCMFSA_2:69; A15: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Initialized_s)_._f_=_(Comput_((P_+*_(if>0_(a,I,J))),(Initialized_s),2))_._f let f be FinSeq-Location ; ::_thesis: (Initialized s) . f = (Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)) . f thus (Initialized s) . f = (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) . f by A10, SCMFSA_2:71 .= (Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)) . f by A13, SCMFSA_2:69 ; ::_thesis: verum end; now__::_thesis:_for_a_being_Int-Location_holds_(Initialized_s)_._a_=_(Comput_((P_+*_(if>0_(a,I,J))),(Initialized_s),2))_._a let a be Int-Location; ::_thesis: (Initialized s) . a = (Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)) . a thus (Initialized s) . a = (Comput ((P +* (if>0 (a,I,J))),(Initialized s),1)) . a by A10, SCMFSA_2:71 .= (Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)) . a by A13, SCMFSA_2:69 ; ::_thesis: verum end; then A16: DataPart (Initialized s) = DataPart (Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)) by A15, SCMFSA_M:2; assume A17: J is_halting_on Initialized s,P ; ::_thesis: IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) then A18: P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) halts_on Initialized s by A12, SCMFSA8A:39; ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on Initialized s,P by A12, A17, SCMFSA8A:37; then A19: ((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA) is_closed_on Initialized s,P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)) by Th6; A20: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2)))) = CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)))))) by EXTPRO_1:4 .= IncAddr ((CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))))),2) by A19, A14, A16, Th8, A9, A1 .= IncAddr ((halt SCM+FSA),2) by A18, EXTPRO_1:def_15 .= halt SCM+FSA by COMPOS_0:4 ; then A21: P +* (if>0 (a,I,J)) halts_on Initialized s by EXTPRO_1:29; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_(LifeSpan_((P_+*_(((J_";"_(Goto_((card_I)_+_1)))_";"_I)_";"_(Stop_SCM+FSA))),(Initialized_s)))_+_2_holds_ CurInstr_((P_+*_(if>0_(a,I,J))),(Comput_((P_+*_(if>0_(a,I,J))),(Initialized_s),l)))_<>_halt_SCM+FSA let l be Element of NAT ; ::_thesis: ( l < (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 implies CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA ) assume A22: l < (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 ; ::_thesis: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA A23: Comput ((P +* (if>0 (a,I,J))),(Initialized s),0) = Initialized s ; percases ( l = 0 or l = 1 or ( l <> 0 & l <> 1 ) ) ; suppose l = 0 ; ::_thesis: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA hence CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) <> halt SCM+FSA by A7, A2, A23, PBOOLE:143; ::_thesis: verum end; suppose l = 1 ; ::_thesis: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) <> halt SCM+FSA hence CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) <> halt SCM+FSA by A11, A4, PBOOLE:143; ::_thesis: verum end; supposeA24: ( l <> 0 & l <> 1 ) ; ::_thesis: not CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),b1))) = halt SCM+FSA assume A25: CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA ; ::_thesis: contradiction consider n being Nat such that A26: l = n + 1 by A24, NAT_1:6; n <> 0 by A24, A26; then consider l2 being Nat such that A27: n = l2 + 1 by NAT_1:6; reconsider l2 = l2 as Element of NAT by ORDINAL1:def_12; InsCode (CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),l2)))) = InsCode (IncAddr ((CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),l2)))),2)) by COMPOS_0:def_9 .= InsCode (CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)),l2)))) by A19, A14, A16, Th8, A9, A1 .= InsCode (CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),(l2 + (1 + 1)))))) by EXTPRO_1:4 .= 0 by A26, A27, A25, COMPOS_1:70 ; then A28: CurInstr ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),l2))) = halt SCM+FSA by SCMFSA_2:95; n + 1 < ((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 1) + 1 by A22, A26; then n < (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 1 by XREAL_1:6; then l2 < LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)) by A27, XREAL_1:6; hence contradiction by A18, A28, EXTPRO_1:def_15; ::_thesis: verum end; end; end; then for l being Element of NAT st CurInstr ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),l))) = halt SCM+FSA holds (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 <= l ; then A29: LifeSpan ((P +* (if>0 (a,I,J))),(Initialized s)) = (LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2 by A20, A21, EXTPRO_1:def_15; A30: DataPart (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) = DataPart (Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))) by A17, A12, EXTPRO_1:23, SCMFSA8A:39 .= DataPart (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))) by A19, A14, A16, Th8, A9, A1 .= DataPart (Comput ((P +* (if>0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2))) by EXTPRO_1:4 .= DataPart (Result ((P +* (if>0 (a,I,J))),(Initialized s))) by A21, A29, EXTPRO_1:23 ; A31: now__::_thesis:_for_x_being_set_st_x_in_dom_(IExec_((if>0_(a,I,J)),P,s))_holds_ (IExec_((if>0_(a,I,J)),P,s))_._x_=_((IExec_((((J_";"_(Goto_((card_I)_+_1)))_";"_I)_";"_(Stop_SCM+FSA)),P,s))_+*_(Start-At_((((card_I)_+_(card_J))_+_3),SCM+FSA)))_._x let x be set ; ::_thesis: ( x in dom (IExec ((if>0 (a,I,J)),P,s)) implies (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 ) A32: dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) = {(IC )} by FUNCOP_1:13; A33: IExec ((if>0 (a,I,J)),P,s) = Result ((P +* (if>0 (a,I,J))),(Initialized s)) by SCMFSA6B:def_1; A34: IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s) = Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)) by SCMFSA6B:def_1; assume A35: x in dom (IExec ((if>0 (a,I,J)),P,s)) ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 percases ( x is Int-Location or x is FinSeq-Location or x = IC ) by A35, SCMFSA_M:1; supposeA36: x is Int-Location ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:56; then A37: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A32, TARSKI:def_1; thus (IExec ((if>0 (a,I,J)),P,s)) . x = (Result ((P +* (if>0 (a,I,J))),(Initialized s))) . x by A33 .= (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) . x by A30, A36, SCMFSA_M:2 .= (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) . x by A34 .= ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A37, FUNCT_4:11 ; ::_thesis: verum end; supposeA38: x is FinSeq-Location ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then x <> IC by SCMFSA_2:57; then A39: not x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A32, TARSKI:def_1; thus (IExec ((if>0 (a,I,J)),P,s)) . x = (Result ((P +* (if>0 (a,I,J))),(Initialized s))) . x by A33 .= (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) . x by A30, A38, SCMFSA_M:2 .= (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) . x by A34 .= ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A39, FUNCT_4:11 ; ::_thesis: verum end; supposeA40: x = IC ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . b1 = ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . b1 then A41: x in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A32, TARSKI:def_1; A42: IC (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) = IC (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) by A34 .= ((card I) + (card J)) + 1 by A12, A17, SCMFSA8A:40 ; thus (IExec ((if>0 (a,I,J)),P,s)) . x = (Result ((P +* (if>0 (a,I,J))),(Initialized s))) . x by A33 .= (Comput ((P +* (if>0 (a,I,J))),(Initialized s),((LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))) + 2))) . x by A21, A29, EXTPRO_1:23 .= IC (Comput ((P +* (if>0 (a,I,J))),(Comput ((P +* (if>0 (a,I,J))),(Initialized s),2)),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s))))) by A40, EXTPRO_1:4 .= (IC (Comput ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s),(LifeSpan ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)))))) + 2 by A19, A14, A16, Th8, A9, A1 .= (IC (Result ((P +* (((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA))),(Initialized s)))) + 2 by A17, A12, EXTPRO_1:23, SCMFSA8A:39 .= (Start-At (((((card I) + (card J)) + 1) + 2),SCM+FSA)) . (IC ) by A42, FUNCOP_1:72 .= ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) . x by A40, A41, FUNCT_4:13 ; ::_thesis: verum end; end; end; dom (IExec ((if>0 (a,I,J)),P,s)) = the carrier of SCM+FSA by PARTFUN1:def_2 .= dom ((IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) by PARTFUN1:def_2 ; hence IExec ((if>0 (a,I,J)),P,s) = (IExec ((((J ";" (Goto ((card I) + 1))) ";" I) ";" (Stop SCM+FSA)),P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A31, FUNCT_1:2 .= ((IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 1),SCM+FSA))) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by A12, A17, SCMFSA8A:41 .= (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by FUNCT_4:114 ; ::_thesis: verum end; theorem Th23: :: SCMFSA8B:23 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if>0 (a,I,J) is parahalting & ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if>0 (a,I,J) is parahalting & ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if>0 (a,I,J) is parahalting & ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a being read-write Int-Location holds ( if>0 (a,I,J) is parahalting & ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) let a be read-write Int-Location; ::_thesis: ( if>0 (a,I,J) is parahalting & ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) A1: I is_halting_on Initialized s,P by SCMFSA7B:19; for s being 0 -started State of SCM+FSA for P being Instruction-Sequence of SCM+FSA st if>0 (a,I,J) c= P holds P halts_on s proof let s be 0 -started State of SCM+FSA; ::_thesis: for P being Instruction-Sequence of SCM+FSA st if>0 (a,I,J) c= P holds P halts_on s Start-At (0,SCM+FSA) c= s by MEMSTR_0:29; then A2: s = Initialize s by FUNCT_4:98; let Q be Instruction-Sequence of SCM+FSA; ::_thesis: ( if>0 (a,I,J) c= Q implies Q halts_on s ) assume A3: if>0 (a,I,J) c= Q ; ::_thesis: Q halts_on s A4: if>0 (a,I,J) c= Q by A3; A5: I is_closed_on s,Q by SCMFSA7B:18; A6: I is_halting_on s,Q by SCMFSA7B:19; A7: J is_halting_on s,Q by SCMFSA7B:19; A8: J is_closed_on s,Q by SCMFSA7B:18; A9: Q +* (if>0 (a,I,J)) = Q by A4, FUNCT_4:98; percases ( s . a > 0 or s . a <= 0 ) ; suppose s . a > 0 ; ::_thesis: Q halts_on s then if>0 (a,I,J) is_halting_on s,Q by A5, A6, Th19; hence Q halts_on s by A2, A9, SCMFSA7B:def_7; ::_thesis: verum end; suppose s . a <= 0 ; ::_thesis: Q halts_on s then if>0 (a,I,J) is_halting_on s,Q by A8, A7, Th21; hence Q halts_on s by A2, A9, SCMFSA7B:def_7; ::_thesis: verum end; end; end; hence if>0 (a,I,J) is parahalting by AMISTD_1:def_11; ::_thesis: ( ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) & ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) ) A10: J is_closed_on Initialized s,P by SCMFSA7B:18; I is_closed_on Initialized s,P by SCMFSA7B:18; hence ( s . a > 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) by A1, Th20; ::_thesis: ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) J is_halting_on Initialized s,P by SCMFSA7B:19; hence ( s . a <= 0 implies IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) ) by Th22, A10; ::_thesis: verum end; theorem Th24: :: SCMFSA8B:24 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a > 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a > 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a > 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a being read-write Int-Location holds ( IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a > 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let a be read-write Int-Location; ::_thesis: ( IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 & ( s . a > 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) hereby ::_thesis: ( ( s . a > 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) percases ( s . a > 0 or s . a <= 0 ) ; suppose s . a > 0 ; ::_thesis: IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 then IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th23; hence IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 by FUNCT_4:113; ::_thesis: verum end; suppose s . a <= 0 ; ::_thesis: IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 then IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th23; hence IC (IExec ((if>0 (a,I,J)),P,s)) = ((card I) + (card J)) + 3 by FUNCT_4:113; ::_thesis: verum end; end; end; hereby ::_thesis: ( s . a <= 0 implies ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) assume s . a > 0 ; ::_thesis: ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) then A1: IExec ((if>0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th23; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f let d be Int-Location; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d not d in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:102; hence (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (I,P,s)) . d by A1, FUNCT_4:11; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f not f in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:103; hence (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (I,P,s)) . f by A1, FUNCT_4:11; ::_thesis: verum end; assume s . a <= 0 ; ::_thesis: ( ( for d being Int-Location holds (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) then A2: IExec ((if>0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by Th23; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f let d be Int-Location; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d not d in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:102; hence (IExec ((if>0 (a,I,J)),P,s)) . d = (IExec (J,P,s)) . d by A2, FUNCT_4:11; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f not f in dom (Start-At ((((card I) + (card J)) + 3),SCM+FSA)) by SCMFSA_2:103; hence (IExec ((if>0 (a,I,J)),P,s)) . f = (IExec (J,P,s)) . f by A2, FUNCT_4:11; ::_thesis: verum end; theorem :: SCMFSA8B:25 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a < 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a < 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a < 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a < 0 & I is_closed_on s,P & I is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) let a be read-write Int-Location; ::_thesis: ( s . a < 0 & I is_closed_on s,P & I is_halting_on s,P implies ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) ) assume A1: s . a < 0 ; ::_thesis: ( not I is_closed_on s,P or not I is_halting_on s,P or ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) ) assume that A2: I is_closed_on s,P and A3: I is_halting_on s,P ; ::_thesis: ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) A4: if>0 (a,J,I) is_halting_on s,P by A1, A2, A3, Th21; if>0 (a,J,I) is_closed_on s,P by A1, A2, A3, Th21; hence ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) by A1, A4, Th15; ::_thesis: verum end; theorem Th26: :: SCMFSA8B:26 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a < 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a < 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a < 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a < 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let a be read-write Int-Location; ::_thesis: ( s . a < 0 & I is_closed_on Initialized s,P & I is_halting_on Initialized s,P implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) assume A1: s . a < 0 ; ::_thesis: ( not I is_closed_on Initialized s,P or not I is_halting_on Initialized s,P or IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) assume that A2: I is_closed_on Initialized s,P and A3: I is_halting_on Initialized s,P ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) A4: (Initialized s) . a <= 0 by A1, SCMFSA_M:37; then A5: if>0 (a,J,I) is_halting_on Initialized s,P by A2, A3, Th21; if>0 (a,J,I) is_closed_on Initialized s,P by A2, A3, A4, Th21; hence IExec ((if<0 (a,I,J)),P,s) = (IExec ((if>0 (a,J,I)),P,s)) +* (Start-At ((((card (if>0 (a,J,I))) + (card J)) + 3),SCM+FSA)) by A1, A5, Th16 .= (IExec ((if>0 (a,J,I)),P,s)) +* (Start-At ((((((card I) + (card J)) + 4) + (card J)) + 3),SCM+FSA)) by Th12 .= ((IExec (I,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) +* (Start-At ((((((card I) + (card J)) + 4) + (card J)) + 3),SCM+FSA)) by A1, A2, A3, Th22 .= (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) by FUNCT_4:114 ; ::_thesis: verum end; theorem :: SCMFSA8B:27 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) by Th13; theorem Th28: :: SCMFSA8B:28 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a = 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a = 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let a be read-write Int-Location; ::_thesis: ( s . a = 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) assume A1: s . a = 0 ; ::_thesis: ( not J is_closed_on Initialized s,P or not J is_halting_on Initialized s,P or IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) assume that A2: J is_closed_on Initialized s,P and A3: J is_halting_on Initialized s,P ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) thus IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At ((((card (if>0 (a,J,I))) + (card J)) + 3),SCM+FSA)) by A1, A2, A3, Th14 .= (IExec (J,P,s)) +* (Start-At ((((((card I) + (card J)) + 4) + (card J)) + 3),SCM+FSA)) by Th12 .= (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ; ::_thesis: verum end; theorem :: SCMFSA8B:29 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a > 0 & J is_closed_on s,P & J is_halting_on s,P holds ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) let a be read-write Int-Location; ::_thesis: ( s . a > 0 & J is_closed_on s,P & J is_halting_on s,P implies ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) ) assume A1: s . a > 0 ; ::_thesis: ( not J is_closed_on s,P or not J is_halting_on s,P or ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) ) assume that A2: J is_closed_on s,P and A3: J is_halting_on s,P ; ::_thesis: ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) A4: if>0 (a,J,I) is_halting_on s,P by A1, A2, A3, Th19; if>0 (a,J,I) is_closed_on s,P by A1, A2, A3, Th19; hence ( if<0 (a,I,J) is_closed_on s,P & if<0 (a,I,J) is_halting_on s,P ) by A1, A4, Th15; ::_thesis: verum end; theorem Th30: :: SCMFSA8B:30 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for a being read-write Int-Location st s . a > 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let I, J be Program of SCM+FSA; ::_thesis: for a being read-write Int-Location st s . a > 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P holds IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) let a be read-write Int-Location; ::_thesis: ( s . a > 0 & J is_closed_on Initialized s,P & J is_halting_on Initialized s,P implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) assume A1: s . a > 0 ; ::_thesis: ( not J is_closed_on Initialized s,P or not J is_halting_on Initialized s,P or IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) assume that A2: J is_closed_on Initialized s,P and A3: J is_halting_on Initialized s,P ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) A4: (Initialized s) . a > 0 by A1, SCMFSA_M:37; then A5: if>0 (a,J,I) is_halting_on Initialized s,P by A2, A3, Th19; if>0 (a,J,I) is_closed_on Initialized s,P by A2, A3, A4, Th19; hence IExec ((if<0 (a,I,J)),P,s) = (IExec ((if>0 (a,J,I)),P,s)) +* (Start-At ((((card (if>0 (a,J,I))) + (card J)) + 3),SCM+FSA)) by A1, A5, Th16 .= (IExec ((if>0 (a,J,I)),P,s)) +* (Start-At ((((((card I) + (card J)) + 4) + (card J)) + 3),SCM+FSA)) by Th12 .= ((IExec (J,P,s)) +* (Start-At ((((card I) + (card J)) + 3),SCM+FSA))) +* (Start-At ((((((card I) + (card J)) + 4) + (card J)) + 3),SCM+FSA)) by A1, A2, A3, Th20 .= (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) by FUNCT_4:114 ; ::_thesis: verum end; theorem :: SCMFSA8B:31 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if<0 (a,I,J) is parahalting & ( s . a < 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) & ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if<0 (a,I,J) is parahalting & ( s . a < 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) & ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a being read-write Int-Location holds ( if<0 (a,I,J) is parahalting & ( s . a < 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) & ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a being read-write Int-Location holds ( if<0 (a,I,J) is parahalting & ( s . a < 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) & ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) ) let a be read-write Int-Location; ::_thesis: ( if<0 (a,I,J) is parahalting & ( s . a < 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) & ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) ) if>0 (a,J,I) is parahalting by Th23; hence if<0 (a,I,J) is parahalting by Th17; ::_thesis: ( ( s . a < 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) & ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) ) A1: I is_closed_on Initialized s,P by SCMFSA7B:18; hereby ::_thesis: ( s . a >= 0 implies IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) ) A2: I is_halting_on Initialized s,P by SCMFSA7B:19; assume s . a < 0 ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) hence IExec ((if<0 (a,I,J)),P,s) = (IExec (I,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) by A2, Th26, A1; ::_thesis: verum end; hereby ::_thesis: verum assume A3: s . a >= 0 ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) A4: J is_halting_on Initialized s,P by SCMFSA7B:19; A5: J is_closed_on Initialized s,P by SCMFSA7B:18; percases ( s . a = 0 or s . a <> 0 ) ; suppose s . a = 0 ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) hence IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) by A4, Th28, A5; ::_thesis: verum end; suppose s . a <> 0 ; ::_thesis: IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) hence IExec ((if<0 (a,I,J)),P,s) = (IExec (J,P,s)) +* (Start-At (((((card I) + (card J)) + (card J)) + 7),SCM+FSA)) by A3, A4, Th30, A5; ::_thesis: verum end; end; end; end; registration let I, J be parahalting Program of SCM+FSA; let a be read-write Int-Location; cluster if=0 (a,I,J) -> parahalting ; correctness coherence if=0 (a,I,J) is parahalting ; by Th17; cluster if>0 (a,I,J) -> parahalting ; correctness coherence if>0 (a,I,J) is parahalting ; by Th23; end; definition let a, b be Int-Location; let I, J be Program of SCM+FSA; func if=0 (a,b,I,J) -> Program of SCM+FSA equals :: SCMFSA8B:def 4 (SubFrom (a,b)) ";" (if=0 (a,I,J)); coherence (SubFrom (a,b)) ";" (if=0 (a,I,J)) is Program of SCM+FSA ; func if>0 (a,b,I,J) -> Program of SCM+FSA equals :: SCMFSA8B:def 5 (SubFrom (a,b)) ";" (if>0 (a,I,J)); coherence (SubFrom (a,b)) ";" (if>0 (a,I,J)) is Program of SCM+FSA ; end; :: deftheorem defines if=0 SCMFSA8B:def_4_:_ for a, b being Int-Location for I, J being Program of SCM+FSA holds if=0 (a,b,I,J) = (SubFrom (a,b)) ";" (if=0 (a,I,J)); :: deftheorem defines if>0 SCMFSA8B:def_5_:_ for a, b being Int-Location for I, J being Program of SCM+FSA holds if>0 (a,b,I,J) = (SubFrom (a,b)) ";" (if>0 (a,I,J)); registration let a be Int-Location; let I, J be Program of SCM+FSA; cluster if=0 (a,I,J) -> non halt-free ; coherence not if=0 (a,I,J) is halt-free ; cluster if>0 (a,I,J) -> non halt-free ; coherence not if>0 (a,I,J) is halt-free ; cluster if<0 (a,I,J) -> non halt-free ; coherence not if<0 (a,I,J) is halt-free ; end; registration let a, b be Int-Location; let I, J be Program of SCM+FSA; cluster if=0 (a,b,I,J) -> non halt-free ; coherence not if=0 (a,b,I,J) is halt-free ; cluster if>0 (a,b,I,J) -> non halt-free ; coherence not if>0 (a,b,I,J) is halt-free ; end; notation let a, b be Int-Location; let I, J be Program of SCM+FSA; synonym if<0 (b,a,I,J) for if>0 (a,b,I,J); end; registration let I, J be parahalting Program of SCM+FSA; let a, b be read-write Int-Location; cluster if=0 (a,b,I,J) -> parahalting ; correctness coherence if=0 (a,b,I,J) is parahalting ; ; cluster if>0 (a,b,I,J) -> parahalting ; correctness coherence if>0 (a,b,I,J) is parahalting ; ; end; theorem :: SCMFSA8B:32 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I being Program of SCM+FSA holds DataPart (Result ((P +* I),(Initialized s))) = DataPart (IExec (I,P,s)) by SCMFSA6B:def_1; theorem Th33: :: SCMFSA8B:33 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I being Program of SCM+FSA holds Result ((P +* I),(Initialized s)) = IExec (I,P,s) by SCMFSA6B:def_1; theorem Th34: :: SCMFSA8B:34 for s1, s2 being State of SCM+FSA for i being Instruction of SCM+FSA for a being Int-Location st ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not i refers a & IC s1 = IC s2 holds ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) proof let s1, s2 be State of SCM+FSA; ::_thesis: for i being Instruction of SCM+FSA for a being Int-Location st ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not i refers a & IC s1 = IC s2 holds ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) let i be Instruction of SCM+FSA; ::_thesis: for a being Int-Location st ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not i refers a & IC s1 = IC s2 holds ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) let a be Int-Location; ::_thesis: ( ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not i refers a & IC s1 = IC s2 implies ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) ) defpred S1[ State of SCM+FSA, State of SCM+FSA] means ( ( for b being Int-Location st a <> b holds $1 . b = $2 . b ) & ( for f being FinSeq-Location holds $1 . f = $2 . f ) ); assume A1: S1[s1,s2] ; ::_thesis: ( i refers a or not IC s1 = IC s2 or ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) ) assume A2: not i refers a ; ::_thesis: ( not IC s1 = IC s2 or ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) ) A3: InsCode i <= 12 by SCMFSA_2:16; A4: now__::_thesis:_for_b_being_Int-Location_st_a_<>_b_holds_ (Exec_(i,s1))_._b_=_(Exec_(i,s2))_._b let b be Int-Location; ::_thesis: ( a <> b implies (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 ) assume A5: a <> b ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 percases ( InsCode i = 0 or InsCode i = 1 or InsCode i = 2 or InsCode i = 3 or InsCode i = 4 or InsCode i = 5 or InsCode i = 6 or InsCode i = 7 or InsCode i = 8 or InsCode i = 9 or InsCode i = 10 or InsCode i = 11 or InsCode i = 12 ) by A3, NAT_1:36; suppose InsCode i = 0 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A6: i = halt SCM+FSA by SCMFSA_2:95; hence (Exec (i,s1)) . b = s1 . b by EXTPRO_1:def_3 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A6, EXTPRO_1:def_3 ; ::_thesis: verum end; suppose InsCode i = 1 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da, db being Int-Location such that A7: i = da := db by SCMFSA_2:30; A8: a <> db by A2, A7, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( b = da or b <> da ) ; supposeA9: b = da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . db by A7, SCMFSA_2:63 .= s2 . db by A1, A8 .= (Exec (i,s2)) . b by A7, A9, SCMFSA_2:63 ; ::_thesis: verum end; supposeA10: b <> da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A7, SCMFSA_2:63 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A7, A10, SCMFSA_2:63 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 2 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da, db being Int-Location such that A11: i = AddTo (da,db) by SCMFSA_2:31; A12: a <> db by A2, A11, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( b = da or b <> da ) ; supposeA13: b = da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = (s1 . b) + (s1 . db) by A11, SCMFSA_2:64 .= (s2 . b) + (s1 . db) by A1, A5 .= (s2 . b) + (s2 . db) by A1, A12 .= (Exec (i,s2)) . b by A11, A13, SCMFSA_2:64 ; ::_thesis: verum end; supposeA14: b <> da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A11, SCMFSA_2:64 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A11, A14, SCMFSA_2:64 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 3 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da, db being Int-Location such that A15: i = SubFrom (da,db) by SCMFSA_2:32; A16: a <> db by A2, A15, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( b = da or b <> da ) ; supposeA17: b = da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = (s1 . b) - (s1 . db) by A15, SCMFSA_2:65 .= (s2 . b) - (s1 . db) by A1, A5 .= (s2 . b) - (s2 . db) by A1, A16 .= (Exec (i,s2)) . b by A15, A17, SCMFSA_2:65 ; ::_thesis: verum end; supposeA18: b <> da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A15, SCMFSA_2:65 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A15, A18, SCMFSA_2:65 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 4 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da, db being Int-Location such that A19: i = MultBy (da,db) by SCMFSA_2:33; A20: a <> db by A2, A19, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( b = da or b <> da ) ; supposeA21: b = da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = (s1 . b) * (s1 . db) by A19, SCMFSA_2:66 .= (s2 . b) * (s1 . db) by A1, A5 .= (s2 . b) * (s2 . db) by A1, A20 .= (Exec (i,s2)) . b by A19, A21, SCMFSA_2:66 ; ::_thesis: verum end; supposeA22: b <> da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A19, SCMFSA_2:66 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A19, A22, SCMFSA_2:66 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 5 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da, db being Int-Location such that A23: i = Divide (da,db) by SCMFSA_2:34; A24: a <> db by A2, A23, SCMFSA7B:def_1; A25: a <> da by A2, A23, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( b = db or ( b = da & b <> db ) or ( b <> da & b <> db ) ) ; supposeA26: b = db ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = (s1 . da) mod (s1 . db) by A23, SCMFSA_2:67 .= (s2 . da) mod (s1 . db) by A1, A25 .= (s2 . da) mod (s2 . db) by A1, A24 .= (Exec (i,s2)) . b by A23, A26, SCMFSA_2:67 ; ::_thesis: verum end; supposeA27: ( b = da & b <> db ) ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = (s1 . da) div (s1 . db) by A23, SCMFSA_2:67 .= (s1 . da) div (s2 . db) by A1, A24 .= (s2 . da) div (s2 . db) by A1, A25 .= (Exec (i,s2)) . b by A23, A27, SCMFSA_2:67 ; ::_thesis: verum end; supposeA28: ( b <> da & b <> db ) ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A23, SCMFSA_2:67 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A23, A28, SCMFSA_2:67 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 6 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A29: ex loc being Element of NAT st i = goto loc by SCMFSA_2:35; hence (Exec (i,s1)) . b = s1 . b by SCMFSA_2:69 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A29, SCMFSA_2:69 ; ::_thesis: verum end; suppose InsCode i = 7 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A30: ex loc being Element of NAT ex da being Int-Location st i = da =0_goto loc by SCMFSA_2:36; hence (Exec (i,s1)) . b = s1 . b by SCMFSA_2:70 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A30, SCMFSA_2:70 ; ::_thesis: verum end; suppose InsCode i = 8 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A31: ex loc being Element of NAT ex da being Int-Location st i = da >0_goto loc by SCMFSA_2:37; hence (Exec (i,s1)) . b = s1 . b by SCMFSA_2:71 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A31, SCMFSA_2:71 ; ::_thesis: verum end; suppose InsCode i = 9 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider db, da being Int-Location, g being FinSeq-Location such that A32: i = da := (g,db) by SCMFSA_2:38; A33: a <> db by A2, A32, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( b = da or b <> da ) ; supposeA34: b = da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b then consider m2 being Element of NAT such that A35: m2 = abs (s2 . db) and A36: (Exec ((da := (g,db)),s2)) . b = (s2 . g) /. m2 by SCMFSA_2:72; consider m1 being Element of NAT such that A37: m1 = abs (s1 . db) and A38: (Exec ((da := (g,db)),s1)) . b = (s1 . g) /. m1 by A34, SCMFSA_2:72; m1 = m2 by A1, A33, A37, A35; hence (Exec (i,s1)) . b = (Exec (i,s2)) . b by A1, A32, A38, A36; ::_thesis: verum end; supposeA39: b <> da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A32, SCMFSA_2:72 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A32, A39, SCMFSA_2:72 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 10 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A40: ex db, da being Int-Location ex g being FinSeq-Location st i = (g,db) := da by SCMFSA_2:39; hence (Exec (i,s1)) . b = s1 . b by SCMFSA_2:73 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A40, SCMFSA_2:73 ; ::_thesis: verum end; suppose InsCode i = 11 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da being Int-Location, g being FinSeq-Location such that A41: i = da :=len g by SCMFSA_2:40; hereby ::_thesis: verum percases ( b = da or b <> da ) ; supposeA42: b = da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = len (s1 . g) by A41, SCMFSA_2:74 .= len (s2 . g) by A1 .= (Exec (i,s2)) . b by A41, A42, SCMFSA_2:74 ; ::_thesis: verum end; supposeA43: b <> da ; ::_thesis: (Exec (i,s1)) . b = (Exec (i,s2)) . b hence (Exec (i,s1)) . b = s1 . b by A41, SCMFSA_2:74 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A41, A43, SCMFSA_2:74 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 12 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A44: ex da being Int-Location ex g being FinSeq-Location st i = g :=<0,...,0> da by SCMFSA_2:41; hence (Exec (i,s1)) . b = s1 . b by SCMFSA_2:75 .= s2 . b by A1, A5 .= (Exec (i,s2)) . b by A44, SCMFSA_2:75 ; ::_thesis: verum end; end; end; assume A45: IC s1 = IC s2 ; ::_thesis: ( ( for b being Int-Location st a <> b holds (Exec (i,s1)) . b = (Exec (i,s2)) . b ) & ( for f being FinSeq-Location holds (Exec (i,s1)) . f = (Exec (i,s2)) . f ) & IC (Exec (i,s1)) = IC (Exec (i,s2)) ) now__::_thesis:_for_f_being_FinSeq-Location_holds_(Exec_(i,s1))_._f_=_(Exec_(i,s2))_._f let f be FinSeq-Location ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 percases ( InsCode i = 0 or InsCode i = 1 or InsCode i = 2 or InsCode i = 3 or InsCode i = 4 or InsCode i = 5 or InsCode i = 6 or InsCode i = 7 or InsCode i = 8 or InsCode i = 9 or InsCode i = 10 or InsCode i = 11 or InsCode i = 12 ) by A3, NAT_1:36; suppose InsCode i = 0 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A46: i = halt SCM+FSA by SCMFSA_2:95; hence (Exec (i,s1)) . f = s1 . f by EXTPRO_1:def_3 .= s2 . f by A1 .= (Exec (i,s2)) . f by A46, EXTPRO_1:def_3 ; ::_thesis: verum end; suppose InsCode i = 1 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A47: ex da, db being Int-Location st i = da := db by SCMFSA_2:30; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:63 .= s2 . f by A1 .= (Exec (i,s2)) . f by A47, SCMFSA_2:63 ; ::_thesis: verum end; suppose InsCode i = 2 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A48: ex da, db being Int-Location st i = AddTo (da,db) by SCMFSA_2:31; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:64 .= s2 . f by A1 .= (Exec (i,s2)) . f by A48, SCMFSA_2:64 ; ::_thesis: verum end; suppose InsCode i = 3 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A49: ex da, db being Int-Location st i = SubFrom (da,db) by SCMFSA_2:32; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:65 .= s2 . f by A1 .= (Exec (i,s2)) . f by A49, SCMFSA_2:65 ; ::_thesis: verum end; suppose InsCode i = 4 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A50: ex da, db being Int-Location st i = MultBy (da,db) by SCMFSA_2:33; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:66 .= s2 . f by A1 .= (Exec (i,s2)) . f by A50, SCMFSA_2:66 ; ::_thesis: verum end; suppose InsCode i = 5 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A51: ex da, db being Int-Location st i = Divide (da,db) by SCMFSA_2:34; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:67 .= s2 . f by A1 .= (Exec (i,s2)) . f by A51, SCMFSA_2:67 ; ::_thesis: verum end; suppose InsCode i = 6 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A52: ex loc being Element of NAT st i = goto loc by SCMFSA_2:35; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:69 .= s2 . f by A1 .= (Exec (i,s2)) . f by A52, SCMFSA_2:69 ; ::_thesis: verum end; suppose InsCode i = 7 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A53: ex loc being Element of NAT ex da being Int-Location st i = da =0_goto loc by SCMFSA_2:36; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:70 .= s2 . f by A1 .= (Exec (i,s2)) . f by A53, SCMFSA_2:70 ; ::_thesis: verum end; suppose InsCode i = 8 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A54: ex loc being Element of NAT ex da being Int-Location st i = da >0_goto loc by SCMFSA_2:37; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:71 .= s2 . f by A1 .= (Exec (i,s2)) . f by A54, SCMFSA_2:71 ; ::_thesis: verum end; suppose InsCode i = 9 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A55: ex db, da being Int-Location ex g being FinSeq-Location st i = da := (g,db) by SCMFSA_2:38; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:72 .= s2 . f by A1 .= (Exec (i,s2)) . f by A55, SCMFSA_2:72 ; ::_thesis: verum end; suppose InsCode i = 10 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider db, da being Int-Location, g being FinSeq-Location such that A56: i = (g,db) := da by SCMFSA_2:39; A57: a <> db by A2, A56, SCMFSA7B:def_1; A58: a <> da by A2, A56, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( f = g or f <> g ) ; supposeA59: f = g ; ::_thesis: (Exec (i,s1)) . f = (Exec (i,s2)) . f A60: s1 . da = s2 . da by A1, A58; consider m2 being Element of NAT such that A61: m2 = abs (s2 . db) and A62: (Exec (((g,db) := da),s2)) . g = (s2 . g) +* (m2,(s2 . da)) by SCMFSA_2:73; consider m1 being Element of NAT such that A63: m1 = abs (s1 . db) and A64: (Exec (((g,db) := da),s1)) . g = (s1 . g) +* (m1,(s1 . da)) by SCMFSA_2:73; m1 = m2 by A1, A57, A63, A61; hence (Exec (i,s1)) . f = (Exec (i,s2)) . f by A1, A56, A59, A64, A62, A60; ::_thesis: verum end; supposeA65: f <> g ; ::_thesis: (Exec (i,s1)) . f = (Exec (i,s2)) . f hence (Exec (i,s1)) . f = s1 . f by A56, SCMFSA_2:73 .= s2 . f by A1 .= (Exec (i,s2)) . f by A56, A65, SCMFSA_2:73 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 11 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then A66: ex da being Int-Location ex g being FinSeq-Location st i = da :=len g by SCMFSA_2:40; hence (Exec (i,s1)) . f = s1 . f by SCMFSA_2:74 .= s2 . f by A1 .= (Exec (i,s2)) . f by A66, SCMFSA_2:74 ; ::_thesis: verum end; suppose InsCode i = 12 ; ::_thesis: (Exec (i,s1)) . b1 = (Exec (i,s2)) . b1 then consider da being Int-Location, g being FinSeq-Location such that A67: i = g :=<0,...,0> da by SCMFSA_2:41; A68: a <> da by A2, A67, SCMFSA7B:def_1; hereby ::_thesis: verum percases ( f = g or f <> g ) ; supposeA69: f = g ; ::_thesis: (Exec (i,s1)) . f = (Exec (i,s2)) . f A70: ex m2 being Element of NAT st ( m2 = abs (s2 . da) & (Exec ((g :=<0,...,0> da),s2)) . g = m2 |-> 0 ) by SCMFSA_2:75; ex m1 being Element of NAT st ( m1 = abs (s1 . da) & (Exec ((g :=<0,...,0> da),s1)) . g = m1 |-> 0 ) by SCMFSA_2:75; hence (Exec (i,s1)) . f = (Exec (i,s2)) . f by A1, A67, A68, A69, A70; ::_thesis: verum end; supposeA71: f <> g ; ::_thesis: (Exec (i,s1)) . f = (Exec (i,s2)) . f hence (Exec (i,s1)) . f = s1 . f by A67, SCMFSA_2:75 .= s2 . f by A1 .= (Exec (i,s2)) . f by A67, A71, SCMFSA_2:75 ; ::_thesis: verum end; end; end; end; end; end; hence S1[ Exec (i,s1), Exec (i,s2)] by A4; ::_thesis: IC (Exec (i,s1)) = IC (Exec (i,s2)) now__::_thesis:_(Exec_(i,s1))_._(IC_)_=_(Exec_(i,s2))_._(IC_) percases ( InsCode i = 0 or InsCode i = 1 or InsCode i = 2 or InsCode i = 3 or InsCode i = 4 or InsCode i = 5 or InsCode i = 6 or InsCode i = 7 or InsCode i = 8 or InsCode i = 9 or InsCode i = 10 or InsCode i = 11 or InsCode i = 12 ) by A3, NAT_1:36; suppose InsCode i = 0 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A72: i = halt SCM+FSA by SCMFSA_2:95; hence (Exec (i,s1)) . (IC ) = IC s1 by EXTPRO_1:def_3 .= (Exec (i,s2)) . (IC ) by A45, A72, EXTPRO_1:def_3 ; ::_thesis: verum end; suppose InsCode i = 1 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A73: ex da, db being Int-Location st i = da := db by SCMFSA_2:30; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:63 .= (Exec (i,s2)) . (IC ) by A45, A73, SCMFSA_2:63 ; ::_thesis: verum end; suppose InsCode i = 2 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A74: ex da, db being Int-Location st i = AddTo (da,db) by SCMFSA_2:31; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:64 .= (Exec (i,s2)) . (IC ) by A45, A74, SCMFSA_2:64 ; ::_thesis: verum end; suppose InsCode i = 3 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A75: ex da, db being Int-Location st i = SubFrom (da,db) by SCMFSA_2:32; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:65 .= (Exec (i,s2)) . (IC ) by A45, A75, SCMFSA_2:65 ; ::_thesis: verum end; suppose InsCode i = 4 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A76: ex da, db being Int-Location st i = MultBy (da,db) by SCMFSA_2:33; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:66 .= (Exec (i,s2)) . (IC ) by A45, A76, SCMFSA_2:66 ; ::_thesis: verum end; suppose InsCode i = 5 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A77: ex da, db being Int-Location st i = Divide (da,db) by SCMFSA_2:34; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:67 .= (Exec (i,s2)) . (IC ) by A45, A77, SCMFSA_2:67 ; ::_thesis: verum end; suppose InsCode i = 6 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then consider loc being Element of NAT such that A78: i = goto loc by SCMFSA_2:35; thus (Exec (i,s1)) . (IC ) = loc by A78, SCMFSA_2:69 .= (Exec (i,s2)) . (IC ) by A78, SCMFSA_2:69 ; ::_thesis: verum end; suppose InsCode i = 7 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then consider loc being Element of NAT , da being Int-Location such that A79: i = da =0_goto loc by SCMFSA_2:36; a <> da by A2, A79, SCMFSA7B:def_1; then A80: s1 . da = s2 . da by A1; hereby ::_thesis: verum percases ( s1 . da = 0 or s1 . da <> 0 ) ; supposeA81: s1 . da = 0 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) hence (Exec (i,s1)) . (IC ) = loc by A79, SCMFSA_2:70 .= (Exec (i,s2)) . (IC ) by A79, A80, A81, SCMFSA_2:70 ; ::_thesis: verum end; supposeA82: s1 . da <> 0 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) hence (Exec (i,s1)) . (IC ) = succ (IC s1) by A79, SCMFSA_2:70 .= (Exec (i,s2)) . (IC ) by A45, A79, A80, A82, SCMFSA_2:70 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 8 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then consider loc being Element of NAT , da being Int-Location such that A83: i = da >0_goto loc by SCMFSA_2:37; a <> da by A2, A83, SCMFSA7B:def_1; then A84: s1 . da = s2 . da by A1; hereby ::_thesis: verum percases ( s1 . da > 0 or s1 . da <= 0 ) ; supposeA85: s1 . da > 0 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) hence (Exec (i,s1)) . (IC ) = loc by A83, SCMFSA_2:71 .= (Exec (i,s2)) . (IC ) by A83, A84, A85, SCMFSA_2:71 ; ::_thesis: verum end; supposeA86: s1 . da <= 0 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) hence (Exec (i,s1)) . (IC ) = succ (IC s1) by A83, SCMFSA_2:71 .= (Exec (i,s2)) . (IC ) by A45, A83, A84, A86, SCMFSA_2:71 ; ::_thesis: verum end; end; end; end; suppose InsCode i = 9 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A87: ex db, da being Int-Location ex g being FinSeq-Location st i = da := (g,db) by SCMFSA_2:38; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:72 .= (Exec (i,s2)) . (IC ) by A45, A87, SCMFSA_2:72 ; ::_thesis: verum end; suppose InsCode i = 10 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A88: ex db, da being Int-Location ex g being FinSeq-Location st i = (g,db) := da by SCMFSA_2:39; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:73 .= (Exec (i,s2)) . (IC ) by A45, A88, SCMFSA_2:73 ; ::_thesis: verum end; suppose InsCode i = 11 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A89: ex da being Int-Location ex g being FinSeq-Location st i = da :=len g by SCMFSA_2:40; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:74 .= (Exec (i,s2)) . (IC ) by A45, A89, SCMFSA_2:74 ; ::_thesis: verum end; suppose InsCode i = 12 ; ::_thesis: (Exec (i,s1)) . (IC ) = (Exec (i,s2)) . (IC ) then A90: ex da being Int-Location ex g being FinSeq-Location st i = g :=<0,...,0> da by SCMFSA_2:41; hence (Exec (i,s1)) . (IC ) = succ (IC s1) by SCMFSA_2:75 .= (Exec (i,s2)) . (IC ) by A45, A90, SCMFSA_2:75 ; ::_thesis: verum end; end; end; hence IC (Exec (i,s1)) = IC (Exec (i,s2)) ; ::_thesis: verum end; theorem Th35: :: SCMFSA8B:35 for P1, P2 being Instruction-Sequence of SCM+FSA for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 holds for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) proof let P1, P2 be Instruction-Sequence of SCM+FSA; ::_thesis: for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 holds for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) let s1, s2 be State of SCM+FSA; ::_thesis: for I being Program of SCM+FSA for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 holds for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) let I be Program of SCM+FSA; ::_thesis: for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 holds for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) let a be Int-Location; ::_thesis: ( not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 implies for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) ) assume A1: not I refers a ; ::_thesis: ( ex b being Int-Location st ( a <> b & not s1 . b = s2 . b ) or ex f being FinSeq-Location st not s1 . f = s2 . f or not I is_closed_on s1,P1 or for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) ) set S2 = Initialize s2; set Q2 = P2 +* I; set S1 = Initialize s1; set Q1 = P1 +* I; A2: I c= P1 +* I by FUNCT_4:25; A3: I c= P2 +* I by FUNCT_4:25; defpred S1[ State of SCM+FSA, State of SCM+FSA] means ( ( for b being Int-Location st a <> b holds $1 . b = $2 . b ) & ( for f being FinSeq-Location holds $1 . f = $2 . f ) ); assume that A4: for b being Int-Location st a <> b holds s1 . b = s2 . b and A5: for f being FinSeq-Location holds s1 . f = s2 . f ; ::_thesis: ( not I is_closed_on s1,P1 or for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) ) A6: IC in dom (Start-At (0,SCM+FSA)) by MEMSTR_0:15; A7: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Initialize_s1)_._f_=_(Initialize_s2)_._f let f be FinSeq-Location ; ::_thesis: (Initialize s1) . f = (Initialize s2) . f A8: not f in dom (Start-At (0,SCM+FSA)) by SCMFSA_2:103; hence (Initialize s1) . f = s1 . f by FUNCT_4:11 .= s2 . f by A5 .= (Initialize s2) . f by A8, FUNCT_4:11 ; ::_thesis: verum end; defpred S2[ Nat] means ( S1[ Comput ((P1 +* I),(Initialize s1),$1), Comput ((P2 +* I),(Initialize s2),$1)] & IC (Comput ((P1 +* I),(Initialize s1),$1)) = IC (Comput ((P2 +* I),(Initialize s2),$1)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),$1))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),$1))) ); A9: IC (Comput ((P1 +* I),(Initialize s1),0)) = IC (Initialize s1) .= IC (Start-At (0,SCM+FSA)) by A6, FUNCT_4:13 .= IC (Initialize s2) by A6, FUNCT_4:13 .= IC (Comput ((P2 +* I),(Initialize s2),0)) ; A10: now__::_thesis:_for_b_being_Int-Location_st_a_<>_b_holds_ (Initialize_s1)_._b_=_(Initialize_s2)_._b let b be Int-Location; ::_thesis: ( a <> b implies (Initialize s1) . b = (Initialize s2) . b ) assume A11: a <> b ; ::_thesis: (Initialize s1) . b = (Initialize s2) . b A12: not b in dom (Start-At (0,SCM+FSA)) by SCMFSA_2:102; hence (Initialize s1) . b = s1 . b by FUNCT_4:11 .= s2 . b by A4, A11 .= (Initialize s2) . b by A12, FUNCT_4:11 ; ::_thesis: verum end; assume A13: I is_closed_on s1,P1 ; ::_thesis: for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) A14: IC (Comput ((P1 +* I),(Initialize s1),0)) in dom I by A13, SCMFSA7B:def_6; A15: for k being Element of NAT st S2[k] holds S2[k + 1] proof let k be Element of NAT ; ::_thesis: ( S2[k] implies S2[k + 1] ) A16: Comput ((P1 +* I),(Initialize s1),(k + 1)) = Following ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) by EXTPRO_1:3 .= Exec ((CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k)))),(Comput ((P1 +* I),(Initialize s1),k))) ; A17: IC (Comput ((P1 +* I),(Initialize s1),k)) in dom I by A13, SCMFSA7B:def_6; A18: Comput ((P2 +* I),(Initialize s2),(k + 1)) = Following ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) by EXTPRO_1:3 .= Exec ((CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k)))),(Comput ((P2 +* I),(Initialize s2),k))) ; A19: (P1 +* I) /. (IC (Comput ((P1 +* I),(Initialize s1),k))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),k))) by PBOOLE:143; CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = I . (IC (Comput ((P1 +* I),(Initialize s1),k))) by A17, A19, A2, GRFUNC_1:2; then CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) in rng I by A17, FUNCT_1:def_3; then A20: not CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) refers a by A1, SCMFSA7B:def_2; assume A21: S2[k] ; ::_thesis: S2[k + 1] hence S1[ Comput ((P1 +* I),(Initialize s1),(k + 1)), Comput ((P2 +* I),(Initialize s2),(k + 1))] by A16, A18, A20, Th34; ::_thesis: ( IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) = IC (Comput ((P2 +* I),(Initialize s2),(k + 1))) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),(k + 1)))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),(k + 1)))) ) thus A22: IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) = IC (Comput ((P2 +* I),(Initialize s2),(k + 1))) by A21, A16, A18, A20, Th34; ::_thesis: CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),(k + 1)))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),(k + 1)))) A23: IC (Comput ((P1 +* I),(Initialize s1),(k + 1))) in dom I by A13, SCMFSA7B:def_6; A24: (P1 +* I) /. (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) by PBOOLE:143; A25: (P2 +* I) /. (IC (Comput ((P2 +* I),(Initialize s2),(k + 1)))) = (P2 +* I) . (IC (Comput ((P2 +* I),(Initialize s2),(k + 1)))) by PBOOLE:143; thus CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),(k + 1)))) = I . (IC (Comput ((P1 +* I),(Initialize s1),(k + 1)))) by A23, A24, A2, GRFUNC_1:2 .= CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),(k + 1)))) by A22, A23, A25, A3, GRFUNC_1:2 ; ::_thesis: verum end; CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),0))) = (P1 +* I) . (IC (Comput ((P1 +* I),(Initialize s1),0))) by PBOOLE:143 .= I . (IC (Comput ((P1 +* I),(Initialize s1),0))) by A14, A2, GRFUNC_1:2 .= (P2 +* I) . (IC (Comput ((P2 +* I),(Initialize s2),0))) by A9, A14, A3, GRFUNC_1:2 .= CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),0))) by PBOOLE:143 ; then A26: S2[ 0 ] by A10, A7, A9; for k being Element of NAT holds S2[k] from NAT_1:sch_1(A26, A15); hence for k being Element of NAT holds ( ( for b being Int-Location st a <> b holds (Comput ((P1 +* I),(Initialize s1),k)) . b = (Comput ((P2 +* I),(Initialize s2),k)) . b ) & ( for f being FinSeq-Location holds (Comput ((P1 +* I),(Initialize s1),k)) . f = (Comput ((P2 +* I),(Initialize s2),k)) . f ) & IC (Comput ((P1 +* I),(Initialize s1),k)) = IC (Comput ((P2 +* I),(Initialize s2),k)) & CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),k))) = CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),k))) ) ; ::_thesis: verum end; theorem :: SCMFSA8B:36 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being Program of SCM+FSA for l being Element of NAT holds ( I is_closed_on s,P & I is_halting_on s,P iff ( I is_closed_on s +* (Start-At (l,SCM+FSA)),P +* I & I is_halting_on s +* (Start-At (l,SCM+FSA)),P +* I ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being Program of SCM+FSA for l being Element of NAT holds ( I is_closed_on s,P & I is_halting_on s,P iff ( I is_closed_on s +* (Start-At (l,SCM+FSA)),P +* I & I is_halting_on s +* (Start-At (l,SCM+FSA)),P +* I ) ) let s be State of SCM+FSA; ::_thesis: for I, J being Program of SCM+FSA for l being Element of NAT holds ( I is_closed_on s,P & I is_halting_on s,P iff ( I is_closed_on s +* (Start-At (l,SCM+FSA)),P +* I & I is_halting_on s +* (Start-At (l,SCM+FSA)),P +* I ) ) let I, J be Program of SCM+FSA; ::_thesis: for l being Element of NAT holds ( I is_closed_on s,P & I is_halting_on s,P iff ( I is_closed_on s +* (Start-At (l,SCM+FSA)),P +* I & I is_halting_on s +* (Start-At (l,SCM+FSA)),P +* I ) ) let l be Element of NAT ; ::_thesis: ( I is_closed_on s,P & I is_halting_on s,P iff ( I is_closed_on s +* (Start-At (l,SCM+FSA)),P +* I & I is_halting_on s +* (Start-At (l,SCM+FSA)),P +* I ) ) DataPart s = DataPart (s +* (Start-At (l,SCM+FSA))) by MEMSTR_0:79; hence ( I is_closed_on s,P & I is_halting_on s,P iff ( I is_closed_on s +* (Start-At (l,SCM+FSA)),P +* I & I is_halting_on s +* (Start-At (l,SCM+FSA)),P +* I ) ) by Th5; ::_thesis: verum end; theorem Th37: :: SCMFSA8B:37 for P1, P2 being Instruction-Sequence of SCM+FSA for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) proof let P1, P2 be Instruction-Sequence of SCM+FSA; ::_thesis: for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) let s1, s2 be State of SCM+FSA; ::_thesis: for I being Program of SCM+FSA for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) let I be Program of SCM+FSA; ::_thesis: for a being Int-Location st not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 & I is_halting_on s1,P1 holds ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) let a be Int-Location; ::_thesis: ( not I refers a & ( for b being Int-Location st a <> b holds s1 . b = s2 . b ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & I is_closed_on s1,P1 & I is_halting_on s1,P1 implies ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) assume A1: not I refers a ; ::_thesis: ( ex b being Int-Location st ( a <> b & not s1 . b = s2 . b ) or ex f being FinSeq-Location st not s1 . f = s2 . f or not I is_closed_on s1,P1 or not I is_halting_on s1,P1 or ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) set S2 = Initialize s2; set Q2 = P2 +* I; set S1 = Initialize s1; set Q1 = P1 +* I; assume A2: for b being Int-Location st a <> b holds s1 . b = s2 . b ; ::_thesis: ( ex f being FinSeq-Location st not s1 . f = s2 . f or not I is_closed_on s1,P1 or not I is_halting_on s1,P1 or ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) assume A3: for f being FinSeq-Location holds s1 . f = s2 . f ; ::_thesis: ( not I is_closed_on s1,P1 or not I is_halting_on s1,P1 or ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) ) assume that A4: I is_closed_on s1,P1 and A5: I is_halting_on s1,P1 ; ::_thesis: ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) A6: now__::_thesis:_for_k_being_Element_of_NAT_holds_IC_(Comput_((P2_+*_I),(Initialize_s2),k))_in_dom_I let k be Element of NAT ; ::_thesis: IC (Comput ((P2 +* I),(Initialize s2),k)) in dom I IC (Comput ((P1 +* I),(Initialize s1),k)) in dom I by A4, SCMFSA7B:def_6; hence IC (Comput ((P2 +* I),(Initialize s2),k)) in dom I by A1, Th35, A4, A3, A2; ::_thesis: verum end; P1 +* I halts_on Initialize s1 by A5, SCMFSA7B:def_7; then consider n being Element of NAT such that A7: CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialize s1),n))) = halt SCM+FSA by EXTPRO_1:29; CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialize s2),n))) = halt SCM+FSA by A1, A7, Th35, A4, A3, A2; then P2 +* I halts_on Initialize s2 by EXTPRO_1:29; hence ( I is_closed_on s2,P2 & I is_halting_on s2,P2 ) by A6, SCMFSA7B:def_6, SCMFSA7B:def_7; ::_thesis: verum end; theorem Th38: :: SCMFSA8B:38 for P1, P2 being Instruction-Sequence of SCM+FSA for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA for a being Int-Location st ( for d being read-write Int-Location st a <> d holds s1 . d = s2 . d ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not I refers a & I is_closed_on Initialized s1,P1 & I is_halting_on Initialized s1,P1 holds ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) proof let P1, P2 be Instruction-Sequence of SCM+FSA; ::_thesis: for s1, s2 being State of SCM+FSA for I being Program of SCM+FSA for a being Int-Location st ( for d being read-write Int-Location st a <> d holds s1 . d = s2 . d ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not I refers a & I is_closed_on Initialized s1,P1 & I is_halting_on Initialized s1,P1 holds ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) let s1, s2 be State of SCM+FSA; ::_thesis: for I being Program of SCM+FSA for a being Int-Location st ( for d being read-write Int-Location st a <> d holds s1 . d = s2 . d ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not I refers a & I is_closed_on Initialized s1,P1 & I is_halting_on Initialized s1,P1 holds ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) let I be Program of SCM+FSA; ::_thesis: for a being Int-Location st ( for d being read-write Int-Location st a <> d holds s1 . d = s2 . d ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not I refers a & I is_closed_on Initialized s1,P1 & I is_halting_on Initialized s1,P1 holds ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) let a be Int-Location; ::_thesis: ( ( for d being read-write Int-Location st a <> d holds s1 . d = s2 . d ) & ( for f being FinSeq-Location holds s1 . f = s2 . f ) & not I refers a & I is_closed_on Initialized s1,P1 & I is_halting_on Initialized s1,P1 implies ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) ) assume that A1: for d being read-write Int-Location st a <> d holds s1 . d = s2 . d and A2: for f being FinSeq-Location holds s1 . f = s2 . f ; ::_thesis: ( I refers a or not I is_closed_on Initialized s1,P1 or not I is_halting_on Initialized s1,P1 or ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) ) A3: now__::_thesis:_for_d_being_Int-Location_st_a_<>_d_holds_ (Initialized_s1)_._d_=_(Initialized_s2)_._d let d be Int-Location; ::_thesis: ( a <> d implies (Initialized s1) . b1 = (Initialized s2) . b1 ) assume A4: a <> d ; ::_thesis: (Initialized s1) . b1 = (Initialized s2) . b1 percases ( d = intloc 0 or d <> intloc 0 ) ; supposeA5: d = intloc 0 ; ::_thesis: (Initialized s1) . b1 = (Initialized s2) . b1 hence (Initialized s1) . d = 1 by SCMFSA_M:9 .= (Initialized s2) . d by A5, SCMFSA_M:9 ; ::_thesis: verum end; suppose d <> intloc 0 ; ::_thesis: (Initialized s1) . b1 = (Initialized s2) . b1 then A6: d is read-write by SCMFSA_M:def_2; hence (Initialized s1) . d = s1 . d by SCMFSA_M:37 .= s2 . d by A1, A4, A6 .= (Initialized s2) . d by A6, SCMFSA_M:37 ; ::_thesis: verum end; end; end; set S1 = Initialized s1; set Q1 = P1 +* I; set S2 = Initialized s2; set Q2 = P2 +* I; assume A7: not I refers a ; ::_thesis: ( not I is_closed_on Initialized s1,P1 or not I is_halting_on Initialized s1,P1 or ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) ) A8: Initialized s2 = Initialize (Initialized s2) by MEMSTR_0:44; assume that A9: I is_closed_on Initialized s1,P1 and A10: I is_halting_on Initialized s1,P1 ; ::_thesis: ( ( for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d ) & ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) A11: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Initialized_s1)_._f_=_(Initialized_s2)_._f let f be FinSeq-Location ; ::_thesis: (Initialized s1) . f = (Initialized s2) . f thus (Initialized s1) . f = s1 . f by SCMFSA_M:37 .= s2 . f by A2 .= (Initialized s2) . f by SCMFSA_M:37 ; ::_thesis: verum end; then I is_halting_on Initialized s2,P2 by A7, A9, A10, A3, Th37; then A12: P2 +* I halts_on Initialized s2 by A8, SCMFSA7B:def_7; A13: Initialized s1 = Initialize (Initialized s1) by MEMSTR_0:44; then A14: P1 +* I halts_on Initialized s1 by A10, SCMFSA7B:def_7; now__::_thesis:_for_l_being_Element_of_NAT_st_l_<_LifeSpan_((P1_+*_I),(Initialized_s1))_holds_ CurInstr_((P2_+*_I),(Comput_((P2_+*_I),(Initialized_s2),l)))_<>_halt_SCM+FSA let l be Element of NAT ; ::_thesis: ( l < LifeSpan ((P1 +* I),(Initialized s1)) implies CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialized s2),l))) <> halt SCM+FSA ) assume l < LifeSpan ((P1 +* I),(Initialized s1)) ; ::_thesis: CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialized s2),l))) <> halt SCM+FSA then CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialized s1),l))) <> halt SCM+FSA by A14, EXTPRO_1:def_15; hence CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialized s2),l))) <> halt SCM+FSA by A7, A9, A3, A11, A13, A8, Th35; ::_thesis: verum end; then A15: for l being Element of NAT st CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialized s2),l))) = halt SCM+FSA holds LifeSpan ((P1 +* I),(Initialized s1)) <= l ; CurInstr ((P2 +* I),(Comput ((P2 +* I),(Initialized s2),(LifeSpan ((P1 +* I),(Initialized s1)))))) = CurInstr ((P1 +* I),(Comput ((P1 +* I),(Initialized s1),(LifeSpan ((P1 +* I),(Initialized s1)))))) by A7, A9, A3, A11, A13, A8, Th35 .= halt SCM+FSA by A14, EXTPRO_1:def_15 ; then A16: LifeSpan ((P1 +* I),(Initialized s1)) = LifeSpan ((P2 +* I),(Initialized s2)) by A12, A15, EXTPRO_1:def_15; then A17: Result ((P2 +* I),(Initialized s2)) = Comput ((P2 +* I),(Initialized s2),(LifeSpan ((P1 +* I),(Initialized s1)))) by A12, EXTPRO_1:23; A18: Result ((P1 +* I),(Initialized s1)) = Comput ((P1 +* I),(Initialized s1),(LifeSpan ((P1 +* I),(Initialized s1)))) by A14, EXTPRO_1:23; A19: Result ((P1 +* I),(Initialized s1)) = IExec (I,P1,s1) by Th33; A20: Result ((P2 +* I),(Initialized s2)) = IExec (I,P2,s2) by Th33; thus for d being Int-Location st a <> d holds (IExec (I,P1,s1)) . d = (IExec (I,P2,s2)) . d by A20, A19, A7, A9, A3, A11, A13, A8, A17, A18, Th35; ::_thesis: ( ( for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f ) & IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) ) thus for f being FinSeq-Location holds (IExec (I,P1,s1)) . f = (IExec (I,P2,s2)) . f by A20, A19, A7, A9, A3, A11, A13, A8, A17, A18, Th35; ::_thesis: IC (IExec (I,P1,s1)) = IC (IExec (I,P2,s2)) thus IC (IExec (I,P1,s1)) = IC (Result ((P1 +* I),(Initialized s1))) by SCMFSA6B:def_1 .= IC (Comput ((P1 +* I),(Initialized s1),(LifeSpan ((P1 +* I),(Initialized s1))))) by A14, EXTPRO_1:23 .= IC (Comput ((P2 +* I),(Initialized s2),(LifeSpan ((P2 +* I),(Initialized s2))))) by A7, A9, A3, A11, A13, A8, A16, Th35 .= IC (Result ((P2 +* I),(Initialized s2))) by A12, EXTPRO_1:23 .= IC (IExec (I,P2,s2)) by SCMFSA6B:def_1 ; ::_thesis: verum end; theorem :: SCMFSA8B:39 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if=0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if=0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if=0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if=0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let a, b be read-write Int-Location; ::_thesis: ( not I refers a & not J refers a implies ( IC (IExec ((if=0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) ) assume that A1: not I refers a and A2: not J refers a ; ::_thesis: ( IC (IExec ((if=0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) reconsider JJ = if=0 (a,I,J) as parahalting Program of SCM+FSA ; reconsider II = Macro (SubFrom (a,b)) as keeping_0 parahalting Program of SCM+FSA ; set i = SubFrom (a,b); set s1 = Exec ((SubFrom (a,b)),(Initialized s)); A3: now__::_thesis:_for_c_being_read-write_Int-Location_st_a_<>_c_holds_ (Exec_((SubFrom_(a,b)),(Initialized_s)))_._c_=_s_._c let c be read-write Int-Location; ::_thesis: ( a <> c implies (Exec ((SubFrom (a,b)),(Initialized s))) . c = s . c ) assume a <> c ; ::_thesis: (Exec ((SubFrom (a,b)),(Initialized s))) . c = s . c hence (Exec ((SubFrom (a,b)),(Initialized s))) . c = (Initialized s) . c by SCMFSA_2:65 .= s . c by SCMFSA_M:37 ; ::_thesis: verum end; IExec ((if=0 (a,b,I,J)),P,s) = IncIC ((IExec (JJ,P,(IExec (II,P,s)))),(card II)) by SCMFSA6B:20; hence IC (IExec ((if=0 (a,b,I,J)),P,s)) = (IC (IExec (JJ,P,(IExec (II,P,s))))) + (card II) by FUNCT_4:113 .= (((card I) + (card J)) + 3) + (card II) by Th18 .= (((card I) + (card J)) + 3) + 2 by COMPOS_1:56 .= ((card I) + (card J)) + 5 ; ::_thesis: ( ( s . a = s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) A4: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Exec_((SubFrom_(a,b)),(Initialized_s)))_._f_=_s_._f let f be FinSeq-Location ; ::_thesis: (Exec ((SubFrom (a,b)),(Initialized s))) . f = s . f thus (Exec ((SubFrom (a,b)),(Initialized s))) . f = (Initialized s) . f by SCMFSA_2:65 .= s . f by SCMFSA_M:37 ; ::_thesis: verum end; hereby ::_thesis: ( s . a <> s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) assume A5: s . a = s . b ; ::_thesis: ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) A6: I is_closed_on Initialized s,P by SCMFSA7B:18; A7: I is_halting_on Initialized s,P by SCMFSA7B:19; A8: (Exec ((SubFrom (a,b)),(Initialized s))) . a = ((Initialized s) . a) - ((Initialized s) . b) by SCMFSA_2:65 .= (s . a) - ((Initialized s) . b) by SCMFSA_M:37 .= (s . a) - (s . b) by SCMFSA_M:37 .= 0 by A5 ; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f let d be Int-Location; ::_thesis: ( a <> d implies (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) assume A9: a <> d ; ::_thesis: (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d thus (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by Th9 .= (IExec (I,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by A8, Th18 .= (IExec (I,P,s)) . d by A1, A3, A4, A6, A7, A9, Th38 ; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f thus (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by Th10 .= (IExec (I,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by A8, Th18 .= (IExec (I,P,s)) . f by A1, A3, A4, A6, A7, Th38 ; ::_thesis: verum end; A10: (Exec ((SubFrom (a,b)),(Initialized s))) . a = ((Initialized s) . a) - ((Initialized s) . b) by SCMFSA_2:65 .= (s . a) - ((Initialized s) . b) by SCMFSA_M:37 .= (s . a) - (s . b) by SCMFSA_M:37 ; A11: J is_halting_on Initialized s,P by SCMFSA7B:19; A12: J is_closed_on Initialized s,P by SCMFSA7B:18; assume s . a <> s . b ; ::_thesis: ( ( for d being Int-Location st a <> d holds (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) then A13: (s . a) + (- (s . b)) <> (s . b) + (- (s . b)) ; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f let d be Int-Location; ::_thesis: ( a <> d implies (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) assume A14: a <> d ; ::_thesis: (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d thus (IExec ((if=0 (a,b,I,J)),P,s)) . d = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by Th9 .= (IExec (J,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by A10, A13, Th18 .= (IExec (J,P,s)) . d by A2, A3, A4, A12, A11, A14, Th38 ; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f thus (IExec ((if=0 (a,b,I,J)),P,s)) . f = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by Th10 .= (IExec (J,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by A10, A13, Th18 .= (IExec (J,P,s)) . f by A2, A3, A4, A12, A11, Th38 ; ::_thesis: verum end; theorem :: SCMFSA8B:40 for P being Instruction-Sequence of SCM+FSA for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if>0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) proof let P be Instruction-Sequence of SCM+FSA; ::_thesis: for s being State of SCM+FSA for I, J being parahalting Program of SCM+FSA for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if>0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let s be State of SCM+FSA; ::_thesis: for I, J being parahalting Program of SCM+FSA for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if>0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let I, J be parahalting Program of SCM+FSA; ::_thesis: for a, b being read-write Int-Location st not I refers a & not J refers a holds ( IC (IExec ((if>0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) let a, b be read-write Int-Location; ::_thesis: ( not I refers a & not J refers a implies ( IC (IExec ((if>0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) ) assume that A1: not I refers a and A2: not J refers a ; ::_thesis: ( IC (IExec ((if>0 (a,b,I,J)),P,s)) = ((card I) + (card J)) + 5 & ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) reconsider JJ = if>0 (a,I,J) as parahalting Program of SCM+FSA ; reconsider II = Macro (SubFrom (a,b)) as keeping_0 parahalting Program of SCM+FSA ; set i = SubFrom (a,b); set s1 = Exec ((SubFrom (a,b)),(Initialized s)); A3: now__::_thesis:_for_c_being_read-write_Int-Location_st_a_<>_c_holds_ (Exec_((SubFrom_(a,b)),(Initialized_s)))_._c_=_s_._c let c be read-write Int-Location; ::_thesis: ( a <> c implies (Exec ((SubFrom (a,b)),(Initialized s))) . c = s . c ) assume a <> c ; ::_thesis: (Exec ((SubFrom (a,b)),(Initialized s))) . c = s . c hence (Exec ((SubFrom (a,b)),(Initialized s))) . c = (Initialized s) . c by SCMFSA_2:65 .= s . c by SCMFSA_M:37 ; ::_thesis: verum end; IExec ((if>0 (a,b,I,J)),P,s) = IncIC ((IExec (JJ,P,(IExec (II,P,s)))),(card II)) by SCMFSA6B:20; hence IC (IExec ((if>0 (a,b,I,J)),P,s)) = (IC (IExec (JJ,P,(IExec (II,P,s))))) + (card II) by FUNCT_4:113 .= (((card I) + (card J)) + 3) + (card II) by Th24 .= (((card I) + (card J)) + 3) + 2 by COMPOS_1:56 .= ((card I) + (card J)) + 5 ; ::_thesis: ( ( s . a > s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) ) & ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) ) A4: now__::_thesis:_for_f_being_FinSeq-Location_holds_(Exec_((SubFrom_(a,b)),(Initialized_s)))_._f_=_s_._f let f be FinSeq-Location ; ::_thesis: (Exec ((SubFrom (a,b)),(Initialized s))) . f = s . f thus (Exec ((SubFrom (a,b)),(Initialized s))) . f = (Initialized s) . f by SCMFSA_2:65 .= s . f by SCMFSA_M:37 ; ::_thesis: verum end; hereby ::_thesis: ( s . a <= s . b implies ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) ) A5: (Exec ((SubFrom (a,b)),(Initialized s))) . a = ((Initialized s) . a) - ((Initialized s) . b) by SCMFSA_2:65 .= (s . a) - ((Initialized s) . b) by SCMFSA_M:37 .= (s . a) - (s . b) by SCMFSA_M:37 ; assume s . a > s . b ; ::_thesis: ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f ) ) then A6: (Exec ((SubFrom (a,b)),(Initialized s))) . a > 0 by A5, XREAL_1:50; A7: I is_halting_on Initialized s,P by SCMFSA7B:19; A8: I is_closed_on Initialized s,P by SCMFSA7B:18; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f let d be Int-Location; ::_thesis: ( a <> d implies (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d ) assume A9: a <> d ; ::_thesis: (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (I,P,s)) . d thus (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by Th9 .= (IExec (I,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by A6, Th24 .= (IExec (I,P,s)) . d by A1, A3, A4, A8, A7, A9, Th38 ; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (I,P,s)) . f thus (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by Th10 .= (IExec (I,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by A6, Th24 .= (IExec (I,P,s)) . f by A1, A3, A4, A8, A7, Th38 ; ::_thesis: verum end; A10: (Exec ((SubFrom (a,b)),(Initialized s))) . a = ((Initialized s) . a) - ((Initialized s) . b) by SCMFSA_2:65 .= (s . a) - ((Initialized s) . b) by SCMFSA_M:37 .= (s . a) - (s . b) by SCMFSA_M:37 ; A11: J is_closed_on Initialized s,P by SCMFSA7B:18; A12: J is_halting_on Initialized s,P by SCMFSA7B:19; assume s . a <= s . b ; ::_thesis: ( ( for d being Int-Location st a <> d holds (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) & ( for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f ) ) then A13: (Exec ((SubFrom (a,b)),(Initialized s))) . a <= 0 by A10, XREAL_1:47; hereby ::_thesis: for f being FinSeq-Location holds (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f let d be Int-Location; ::_thesis: ( a <> d implies (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d ) assume A14: a <> d ; ::_thesis: (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (J,P,s)) . d thus (IExec ((if>0 (a,b,I,J)),P,s)) . d = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by Th9 .= (IExec (J,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . d by A13, Th24 .= (IExec (J,P,s)) . d by A2, A3, A4, A11, A12, A14, Th38 ; ::_thesis: verum end; let f be FinSeq-Location ; ::_thesis: (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (J,P,s)) . f thus (IExec ((if>0 (a,b,I,J)),P,s)) . f = (IExec (JJ,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by Th10 .= (IExec (J,P,(Exec ((SubFrom (a,b)),(Initialized s))))) . f by A13, Th24 .= (IExec (J,P,s)) . f by A2, A3, A4, A11, A12, Th38 ; ::_thesis: verum end;