:: Computation of Two Consecutive Program Blocks for SCMPDS
:: by JingChao Chen
::
:: Received June 15, 1999
:: Copyright (c) 1999-2011 Association of Mizar Users


begin

theorem :: SCMPDS_5:1
canceled;

Lm1: card (Stop SCMPDS) = 1
by COMPOS_1:46;

theorem Th2: :: SCMPDS_5:2
for x being set
for i being Instruction of SCMPDS holds
( x in dom (Load i) iff x = 0 )
proof end;

theorem Th3: :: SCMPDS_5:3
for loc being Element of NAT
for I being Program of SCMPDS st loc in dom (stop I) & (stop I) . loc <> halt SCMPDS holds
loc in dom I
proof end;

theorem :: SCMPDS_5:4
for i being Instruction of SCMPDS holds
( dom (Load i) = {0} & (Load i) . 0 = i ) by FUNCOP_1:19, FUNCOP_1:87;

theorem Th5: :: SCMPDS_5:5
for i being Instruction of SCMPDS holds 0 in dom (Load i)
proof end;

theorem Th6: :: SCMPDS_5:6
for i being Instruction of SCMPDS holds card (Load i) = 1
proof end;

theorem :: SCMPDS_5:7
for I being Program of SCMPDS holds card (stop I) = (card I) + 1 by Lm1, AFINSQ_1:20;

theorem Th8: :: SCMPDS_5:8
for i being Instruction of SCMPDS holds card (stop (Load i)) = 2
proof end;

theorem Th9: :: SCMPDS_5:9
for i being Instruction of SCMPDS holds
( 0 in dom (stop (Load i)) & 1 in dom (stop (Load i)) )
proof end;

theorem Th10: :: SCMPDS_5:10
for i being Instruction of SCMPDS holds
( (stop (Load i)) . 0 = i & (stop (Load i)) . 1 = halt SCMPDS )
proof end;

theorem Th11: :: SCMPDS_5:11
for x being set
for i being Instruction of SCMPDS holds
( x in dom (stop (Load i)) iff ( x = 0 or x = 1 ) )
proof end;

theorem :: SCMPDS_5:12
for i being Instruction of SCMPDS holds dom (stop (Load i)) = {0,1}
proof end;

theorem :: SCMPDS_5:13
canceled;

theorem :: SCMPDS_5:14
canceled;

theorem Th15: :: SCMPDS_5:15
for I, J being Program of SCMPDS holds I c= stop (I ';' J)
proof end;

theorem Th16: :: SCMPDS_5:16
for I, J being Program of SCMPDS holds dom (stop I) c= dom (stop (I ';' J))
proof end;

theorem Th17: :: SCMPDS_5:17
for I, J being Program of SCMPDS holds (stop I) +* (stop (I ';' J)) = stop (I ';' J)
proof end;

theorem :: SCMPDS_5:18
canceled;

set SA0 = Start-At (0,SCMPDS);

theorem Th19: :: SCMPDS_5:19
for a being Int_position
for s being State of SCMPDS holds (Initialize s) . a = s . a
proof end;

theorem Th20: :: SCMPDS_5:20
for P1, P2 being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s1, s2 being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS st stop I c= P1 & stop I c= P2 & NPP s1 = NPP s2 holds
for k being Element of NAT holds
( NPP (Comput (P1,s1,k)) = NPP (Comput (P2,s2,k)) & CurInstr (P1,(Comput (P1,s1,k))) = CurInstr (P2,(Comput (P2,s2,k))) )
proof end;

theorem Th21: :: SCMPDS_5:21
for P1, P2 being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s1, s2 being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS st stop I c= P1 & stop I c= P2 & NPP s1 = NPP s2 holds
( LifeSpan (P1,s1) = LifeSpan (P2,s2) & NPP (Result (P1,s1)) = NPP (Result (P2,s2)) )
proof end;

theorem Th22: :: SCMPDS_5:22
for s being State of SCMPDS
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for I being Program of SCMPDS holds IC (IExec (I,P,s)) = IC (Result ((P +* (stop I)),(Initialize s)))
proof end;

theorem Th23: :: SCMPDS_5:23
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS
for J being Program of SCMPDS st stop I c= P holds
for m being Element of NAT st m <= LifeSpan (P,s) holds
NPP (Comput (P,s,m)) = NPP (Comput ((P +* (I ';' J)),s,m))
proof end;

theorem Th24: :: SCMPDS_5:24
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS
for J being Program of SCMPDS st stop I c= P holds
for m being Element of NAT st m <= LifeSpan (P,s) holds
NPP (Comput (P,s,m)) = NPP (Comput ((P +* (stop (I ';' J))),s,m))
proof end;

Lm2: Load ((DataLoc (0,0)) := 0) is parahalting
proof end;

begin

definition
let i be Instruction of SCMPDS;
attr i is No-StopCode means :Def1: :: SCMPDS_5:def 1
i <> halt SCMPDS;
end;

:: deftheorem Def1 defines No-StopCode SCMPDS_5:def 1 :
for i being Instruction of SCMPDS holds
( i is No-StopCode iff i <> halt SCMPDS );

definition
let i be Instruction of SCMPDS;
attr i is parahalting means :Def2: :: SCMPDS_5:def 2
Load i is parahalting ;
end;

:: deftheorem Def2 defines parahalting SCMPDS_5:def 2 :
for i being Instruction of SCMPDS holds
( i is parahalting iff Load i is parahalting );

registration
cluster shiftable No-StopCode parahalting Element of the Instructions of SCMPDS;
existence
ex b1 being Instruction of SCMPDS st
( b1 is No-StopCode & b1 is shiftable & b1 is parahalting )
proof end;
end;

theorem :: SCMPDS_5:25
for k1 being Integer st k1 <> 0 holds
goto k1 is No-StopCode
proof end;

registration
let a be Int_position ;
cluster return a -> No-StopCode ;
coherence
return a is No-StopCode
proof end;
end;

registration
let a be Int_position ;
let k1 be Integer;
cluster a := k1 -> No-StopCode ;
coherence
a := k1 is No-StopCode
proof end;
cluster saveIC (a,k1) -> No-StopCode ;
coherence
saveIC (a,k1) is No-StopCode
proof end;
end;

registration
let a be Int_position ;
let k1, k2 be Integer;
cluster (a,k1) <>0_goto k2 -> No-StopCode ;
coherence
(a,k1) <>0_goto k2 is No-StopCode
proof end;
cluster (a,k1) <=0_goto k2 -> No-StopCode ;
coherence
(a,k1) <=0_goto k2 is No-StopCode
proof end;
cluster (a,k1) >=0_goto k2 -> No-StopCode ;
coherence
(a,k1) >=0_goto k2 is No-StopCode
proof end;
cluster (a,k1) := k2 -> No-StopCode ;
coherence
(a,k1) := k2 is No-StopCode
proof end;
end;

registration
let a be Int_position ;
let k1, k2 be Integer;
cluster AddTo (a,k1,k2) -> No-StopCode ;
coherence
AddTo (a,k1,k2) is No-StopCode
proof end;
end;

registration
let a, b be Int_position ;
let k1, k2 be Integer;
cluster AddTo (a,k1,b,k2) -> No-StopCode ;
coherence
AddTo (a,k1,b,k2) is No-StopCode
proof end;
cluster SubFrom (a,k1,b,k2) -> No-StopCode ;
coherence
SubFrom (a,k1,b,k2) is No-StopCode
proof end;
cluster MultBy (a,k1,b,k2) -> No-StopCode ;
coherence
MultBy (a,k1,b,k2) is No-StopCode
proof end;
cluster Divide (a,k1,b,k2) -> No-StopCode ;
coherence
Divide (a,k1,b,k2) is No-StopCode
proof end;
cluster (a,k1) := (b,k2) -> No-StopCode ;
coherence
(a,k1) := (b,k2) is No-StopCode
proof end;
end;

registration
cluster halt SCMPDS -> parahalting ;
coherence
halt SCMPDS is parahalting
proof end;
end;

registration
let i be parahalting Instruction of SCMPDS;
cluster Load -> parahalting Program of SCMPDS;
coherence
for b1 being Program of SCMPDS st b1 = Load i holds
b1 is parahalting
by Def2;
end;

Lm3: for i being Instruction of SCMPDS st ( for s being State of SCMPDS holds (Exec (i,s)) . (IC ) = succ (IC s) ) holds
Load i is parahalting
proof end;

registration
let a be Int_position ;
let k1 be Integer;
cluster a := k1 -> parahalting ;
coherence
a := k1 is parahalting
proof end;
end;

registration
let a be Int_position ;
let k1, k2 be Integer;
cluster (a,k1) := k2 -> parahalting ;
coherence
(a,k1) := k2 is parahalting
proof end;
cluster AddTo (a,k1,k2) -> parahalting ;
coherence
AddTo (a,k1,k2) is parahalting
proof end;
end;

registration
let a, b be Int_position ;
let k1, k2 be Integer;
cluster AddTo (a,k1,b,k2) -> parahalting ;
coherence
AddTo (a,k1,b,k2) is parahalting
proof end;
cluster SubFrom (a,k1,b,k2) -> parahalting ;
coherence
SubFrom (a,k1,b,k2) is parahalting
proof end;
cluster MultBy (a,k1,b,k2) -> parahalting ;
coherence
MultBy (a,k1,b,k2) is parahalting
proof end;
cluster Divide (a,k1,b,k2) -> parahalting ;
coherence
Divide (a,k1,b,k2) is parahalting
proof end;
cluster (a,k1) := (b,k2) -> parahalting ;
coherence
(a,k1) := (b,k2) is parahalting
proof end;
end;

theorem Th26: :: SCMPDS_5:26
for i being Instruction of SCMPDS st InsCode i = 1 holds
not i is parahalting
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty stored-program definite AMI-Struct of N;
let IT be NAT -defined the Instructions of S -valued Function;
redefine attr IT is halt-free means :Def3: :: SCMPDS_5:def 3
for x being Element of NAT st x in dom IT holds
IT . x <> halt S;
compatibility
( IT is halt-free iff for x being Element of NAT st x in dom IT holds
IT . x <> halt S )
proof end;
end;

:: deftheorem Def3 defines halt-free SCMPDS_5:def 3 :
for N being with_non-empty_elements set
for S being non empty stored-program definite AMI-Struct of N
for IT being NAT -defined the Instructions of b2 -valued Function holds
( IT is halt-free iff for x being Element of NAT st x in dom IT holds
IT . x <> halt S );

registration
cluster non empty V11() Relation-like NAT -defined the carrier of SCMPDS -defined the Instructions of SCMPDS -valued Function-like the Object-Kind of SCMPDS -compatible V32() V68() halt-free parahalting shiftable set ;
existence
ex b1 being Program of SCMPDS st
( b1 is parahalting & b1 is shiftable & b1 is halt-free )
proof end;
end;

registration
let I, J be halt-free Program of SCMPDS;
cluster I ';' J -> halt-free ;
coherence
I ';' J is halt-free
proof end;
end;

registration
let i be No-StopCode Instruction of SCMPDS;
cluster Load -> halt-free ;
coherence
Load i is halt-free
proof end;
end;

registration
let i be No-StopCode Instruction of SCMPDS;
let J be halt-free Program of SCMPDS;
cluster i ';' J -> halt-free ;
coherence
i ';' J is halt-free
;
end;

registration
let I be halt-free Program of SCMPDS;
let j be No-StopCode Instruction of SCMPDS;
cluster I ';' j -> halt-free ;
coherence
I ';' j is halt-free
;
end;

registration
let i, j be No-StopCode Instruction of SCMPDS;
cluster i ';' j -> halt-free ;
coherence
i ';' j is halt-free
;
end;

theorem Th27: :: SCMPDS_5:27
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS st stop I c= P holds
IC (Comput (P,s,(LifeSpan ((P +* (stop I)),s)))) = card I
proof end;

theorem Th28: :: SCMPDS_5:28
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS
for k being Element of NAT st k < LifeSpan ((P +* (stop I)),s) holds
IC (Comput ((P +* (stop I)),s,k)) in dom I
proof end;

theorem Th29: :: SCMPDS_5:29
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS
for k being Element of NAT st I c= P & k <= LifeSpan ((P +* (stop I)),s) holds
NPP (Comput (P,s,k)) = NPP (Comput ((P +* (stop I)),s,k))
proof end;

theorem Th30: :: SCMPDS_5:30
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS st I c= P holds
IC (Comput (P,s,(LifeSpan ((P +* (stop I)),s)))) = card I
proof end;

theorem Th31: :: SCMPDS_5:31
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS holds
( not I c= P or CurInstr (P,(Comput (P,s,(LifeSpan ((P +* (stop I)),s))))) = halt SCMPDS or IC (Comput (P,s,(LifeSpan ((P +* (stop I)),s)))) = card I )
proof end;

theorem Th32: :: SCMPDS_5:32
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS
for k being Element of NAT st I c= P & k < LifeSpan ((P +* (stop I)),s) holds
CurInstr (P,(Comput (P,s,k))) <> halt SCMPDS
proof end;

theorem Th33: :: SCMPDS_5:33
for s being State of SCMPDS
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for I being parahalting Program of SCMPDS
for J being Program of SCMPDS
for k being Element of NAT st k <= LifeSpan ((P +* (stop I)),(Initialize s)) holds
NPP (Comput ((P +* (stop I)),(Initialize s),k)) = NPP (Comput ((P +* (I ';' J)),(Initialize s),k))
proof end;

theorem Th34: :: SCMPDS_5:34
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being parahalting Program of SCMPDS
for J being Program of SCMPDS
for k being Element of NAT st k <= LifeSpan ((P +* (stop I)),s) holds
NPP (Comput ((P +* (stop I)),s,k)) = NPP (Comput ((P +* (stop (I ';' J))),s,k))
proof end;

registration
let I be parahalting Program of SCMPDS;
let J be parahalting shiftable Program of SCMPDS;
cluster I ';' J -> parahalting Program of SCMPDS;
coherence
for b1 being Program of SCMPDS st b1 = I ';' J holds
b1 is parahalting
proof end;
end;

registration
let i be parahalting Instruction of SCMPDS;
let J be parahalting shiftable Program of SCMPDS;
cluster i ';' J -> parahalting ;
coherence
i ';' J is parahalting
;
end;

registration
let I be parahalting Program of SCMPDS;
let j be shiftable parahalting Instruction of SCMPDS;
cluster I ';' j -> parahalting ;
coherence
I ';' j is parahalting
;
end;

registration
let i be parahalting Instruction of SCMPDS;
let j be shiftable parahalting Instruction of SCMPDS;
cluster i ';' j -> parahalting ;
coherence
i ';' j is parahalting
;
end;

theorem Th35: :: SCMPDS_5:35
for m, n being Element of NAT
for P, P1 being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being State of SCMPDS
for s1 being 0 -started State of SCMPDS
for J being parahalting shiftable Program of SCMPDS st stop J c= P & s = Comput ((P1 +* (stop J)),s1,m) holds
Exec ((CurInstr (P,s)),(s +* (Start-At (((IC s) + n),SCMPDS)))) = IncIC ((Following (P,s)),n)
proof end;

begin

theorem :: SCMPDS_5:36
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS
for J being parahalting shiftable Program of SCMPDS
for k being Element of NAT st stop (I ';' J) c= P holds
NPP (IncIC ((Comput (((P +* (stop I)) +* (stop J)),(Initialize (Result ((P +* (stop I)),s))),k)),(card I))) = NPP (Comput ((P +* (stop (I ';' J))),s,((LifeSpan ((P +* (stop I)),s)) + k)))
proof end;

Lm4: for P, P1 being the Instructions of SCMPDS -valued ManySortedSet of NAT
for I being halt-free parahalting Program of SCMPDS
for J being parahalting shiftable Program of SCMPDS
for s, s being 0 -started State of SCMPDS st stop (I ';' J) c= P & P1 = P +* (stop I) holds
( IC (Comput (P,s,(LifeSpan (P1,s)))) = card I & DataPart (Comput (P,s,(LifeSpan (P1,s)))) = DataPart (Initialize (Comput (P1,s,(LifeSpan (P1,s))))) & Shift ((stop J),(card I)) c= P & LifeSpan (P,s) = (LifeSpan (P1,s)) + (LifeSpan ((P1 +* (stop J)),(Initialize (Result (P1,s))))) )
proof end;

theorem Th37: :: SCMPDS_5:37
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS
for J being parahalting shiftable Program of SCMPDS holds LifeSpan ((P +* (stop (I ';' J))),s) = (LifeSpan ((P +* (stop I)),s)) + (LifeSpan (((P +* (stop I)) +* (stop J)),(Initialize (Result ((P +* (stop I)),s)))))
proof end;

theorem Th38: :: SCMPDS_5:38
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS
for J being parahalting shiftable Program of SCMPDS holds IExec ((I ';' J),P,s) = IncIC ((IExec (J,P,(IExec (I,P,s)))),(card I))
proof end;

theorem :: SCMPDS_5:39
for a being Int_position
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS
for J being parahalting shiftable Program of SCMPDS holds (IExec ((I ';' J),P,s)) . a = (IExec (J,P,(IExec (I,P,s)))) . a
proof end;

begin

theorem Th40: :: SCMPDS_5:40
for a being Int_position
for s being State of SCMPDS
for loc being Element of NAT holds
( IC (Initialize s) = 0 & (Initialize s) . a = s . a & (Initialize s) . loc = s . loc )
proof end;

theorem :: SCMPDS_5:41
for s1, s2 being State of SCMPDS holds
( NPP s1 = NPP s2 iff s1 | (SCM-Data-Loc \/ {(IC )}) = s2 | (SCM-Data-Loc \/ {(IC )}) )
proof end;

theorem :: SCMPDS_5:42
canceled;

theorem Th43: :: SCMPDS_5:43
for i being Instruction of SCMPDS
for s1, s2 being State of SCMPDS st DataPart s1 = DataPart s2 & InsCode i <> 3 holds
DataPart (Exec (i,s1)) = DataPart (Exec (i,s2))
proof end;

theorem Th44: :: SCMPDS_5:44
for s1, s2 being State of SCMPDS
for i being shiftable Instruction of SCMPDS st DataPart s1 = DataPart s2 holds
DataPart (Exec (i,s1)) = DataPart (Exec (i,s2))
proof end;

theorem Th45: :: SCMPDS_5:45
for s being State of SCMPDS
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for i being parahalting Instruction of SCMPDS holds Exec (i,(Initialize s)) = IExec ((Load i),P,s)
proof end;

theorem Th46: :: SCMPDS_5:46
for a being Int_position
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for I being halt-free parahalting Program of SCMPDS
for j being shiftable parahalting Instruction of SCMPDS holds (IExec ((I ';' j),P,s)) . a = (Exec (j,(IExec (I,P,s)))) . a
proof end;

theorem :: SCMPDS_5:47
for a being Int_position
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT
for s being 0 -started State of SCMPDS
for i being No-StopCode parahalting Instruction of SCMPDS
for j being shiftable parahalting Instruction of SCMPDS holds (IExec ((i ';' j),P,s)) . a = (Exec (j,(Exec (i,(Initialize s))))) . a
proof end;

theorem :: SCMPDS_5:48
for s being State of SCMPDS
for I being Program of SCMPDS
for P being the Instructions of SCMPDS -valued ManySortedSet of NAT holds IExec (I,P,(Initialize s)) = IExec (I,P,s)
proof end;