:: Initialization Halting Concepts and Their Basic Properties of SCM+FSA
:: by JingChao Chen and Yatsuka Nakamura
::
:: Received June 17, 1998
:: Copyright (c) 1998 Association of Mizar Users


begin

set SA0 = Start-At (0,SCM+FSA);

definition
let I be Program of SCM+FSA;
attr I is InitClosed means :Def1: :: SCM_HALT:def 1
for s being State of SCM+FSA
for P being the Instructions of SCM+FSA -valued ManySortedSet of NAT st I c= P holds
for n being Element of NAT st Initialized I c= s holds
IC (Comput (P,s,n)) in dom I;
attr I is InitHalting means :Def2: :: SCM_HALT:def 2
Initialized I is halting ;
attr I is keepInt0_1 means :Def3: :: SCM_HALT:def 3
for s being State of SCM+FSA st Initialized I c= s holds
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT st I c= p holds
for k being Element of NAT holds (Comput (p,s,k)) . (intloc 0) = 1;
end;

:: deftheorem Def1 defines InitClosed SCM_HALT:def 1 :
for I being Program of SCM+FSA holds
( I is InitClosed iff for s being State of SCM+FSA
for P being the Instructions of SCM+FSA -valued ManySortedSet of NAT st I c= P holds
for n being Element of NAT st Initialized I c= s holds
IC (Comput (P,s,n)) in dom I );

:: deftheorem Def2 defines InitHalting SCM_HALT:def 2 :
for I being Program of SCM+FSA holds
( I is InitHalting iff Initialized I is halting );

:: deftheorem Def3 defines keepInt0_1 SCM_HALT:def 3 :
for I being Program of SCM+FSA holds
( I is keepInt0_1 iff for s being State of SCM+FSA st Initialized I c= s holds
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT st I c= p holds
for k being Element of NAT holds (Comput (p,s,k)) . (intloc 0) = 1 );

theorem Th1: :: SCM_HALT:1
for x being set
for i, m, n being Element of NAT holds
( not x in dom (((intloc i) .--> m) +* (Start-At (n,SCM+FSA))) or x = intloc i or x = IC )
proof end;

theorem Th2: :: SCM_HALT:2
for I being Program of SCM+FSA
for i, m, n being Element of NAT holds dom I misses dom (((intloc i) .--> m) +* (Start-At (n,SCM+FSA)))
proof end;

set iS = Initialize ((intloc 0) .--> 1);

theorem :: SCM_HALT:3
canceled;

theorem Th4: :: SCM_HALT:4
Macro (halt SCM+FSA) is InitHalting
proof end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() V162() InitHalting set ;
existence
ex b1 being Program of SCM+FSA st b1 is InitHalting
by Th4;
end;

theorem Th5: :: SCM_HALT:5
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting Program of SCM+FSA st Initialized I c= s & I c= p holds
p halts_on s
proof end;

theorem Th6: :: SCM_HALT:6
for I being Program of SCM+FSA holds Initialize I c= Initialized I
proof end;

theorem Th7: :: SCM_HALT:7
for I being Program of SCM+FSA
for s being State of SCM+FSA st Initialized I c= s holds
s . (intloc 0) = 1
proof end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() paraclosed -> InitClosed set ;
coherence
for b1 being Program of SCM+FSA st b1 is paraclosed holds
b1 is InitClosed
proof end;
end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() parahalting -> InitHalting set ;
coherence
for b1 being Program of SCM+FSA st b1 is parahalting holds
b1 is InitHalting
proof end;
end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() InitHalting -> InitClosed set ;
coherence
for b1 being Program of SCM+FSA st b1 is InitHalting holds
b1 is InitClosed
proof end;
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() keepInt0_1 -> InitClosed set ;
coherence
for b1 being Program of SCM+FSA st b1 is keepInt0_1 holds
b1 is InitClosed
proof end;
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() keeping_0 -> keepInt0_1 set ;
coherence
for b1 being Program of SCM+FSA st b1 is keeping_0 holds
b1 is keepInt0_1
proof end;
end;

theorem :: SCM_HALT:8
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting Program of SCM+FSA
for a being read-write Int-Location st not a in UsedIntLoc I holds
(IExec (I,p,s)) . a = s . a
proof end;

theorem :: SCM_HALT:9
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting Program of SCM+FSA
for f being FinSeq-Location st not f in UsedInt*Loc I holds
(IExec (I,p,s)) . f = s . f
proof end;

registration
let I be InitHalting Program of SCM+FSA;
cluster Initialized I -> halting ;
coherence
Initialized I is halting
by Def2;
end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() InitHalting -> set ;
coherence
for b1 being Program of SCM+FSA st b1 is InitHalting holds
not b1 is empty
;
end;

theorem :: SCM_HALT:10
canceled;

theorem :: SCM_HALT:11
canceled;

theorem Th12: :: SCM_HALT:12
for s1, s2 being State of SCM+FSA
for p1, p2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for J being InitHalting Program of SCM+FSA st Initialized J c= s1 & J c= p1 holds
for n being Element of NAT st Reloc (J,n) c= s2 & Reloc (J,n) c= p2 & IC s2 = n & DataPart s1 = DataPart s2 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 end;

theorem :: SCM_HALT:13
for I being Program of SCM+FSA
for s being State of SCM+FSA st Initialized I c= s holds
I c= s
proof end;

theorem Th14: :: SCM_HALT:14
for s1, s2 being State of SCM+FSA
for p1, p2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting Program of SCM+FSA st Initialized I c= s1 & Initialized I c= s2 & I c= p1 & I c= p2 & s1,s2 equal_outside NAT holds
for k being Element of NAT holds
( Comput (p1,s1,k), Comput (p2,s2,k) equal_outside NAT & CurInstr (p1,(Comput (p1,s1,k))) = CurInstr (p2,(Comput (p2,s2,k))) )
proof end;

theorem Th15: :: SCM_HALT:15
for s1, s2 being State of SCM+FSA
for p1, p2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting Program of SCM+FSA st Initialized I c= s1 & Initialized I c= s2 & I c= p1 & I c= p2 & s1,s2 equal_outside NAT holds
( LifeSpan (p1,s1) = LifeSpan (p2,s2) & Result (p1,s1), Result (p2,s2) equal_outside NAT )
proof end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() keeping_0 V162() InitHalting set ;
existence
ex b1 being Program of SCM+FSA st
( b1 is keeping_0 & b1 is InitHalting )
proof end;
end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() V162() InitHalting keepInt0_1 set ;
existence
ex b1 being Program of SCM+FSA st
( b1 is keepInt0_1 & b1 is InitHalting )
proof end;
end;

theorem :: SCM_HALT:16
canceled;

theorem Th17: :: SCM_HALT:17
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA holds (IExec (I,p,s)) . (intloc 0) = 1
proof end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() V162() InitClosed set ;
existence
ex b1 being Program of SCM+FSA st b1 is InitClosed
proof end;
end;

theorem Th18: :: SCM_HALT:18
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitClosed Program of SCM+FSA
for J being Program of SCM+FSA st Initialized I c= s & I c= p & p halts_on s holds
for m being Element of NAT st m <= LifeSpan (p,s) holds
Comput (p,s,m), Comput ((p +* (I ';' J)),(s +* (I ';' J)),m) equal_outside NAT
proof end;

theorem Th19: :: SCM_HALT:19
for I being Program of SCM+FSA
for s being State of SCM+FSA
for i, m, n being Element of NAT holds (s +* I) +* (((intloc i) .--> m) +* (Start-At (n,SCM+FSA))) = (s +* (((intloc i) .--> m) +* (Start-At (n,SCM+FSA)))) +* I
proof end;

theorem Th20: :: SCM_HALT:20
for I being Program of SCM+FSA
for s being State of SCM+FSA st Initialize ((intloc 0) .--> 1) c= s holds
( Initialized I c= s +* (I +* (Initialize ((intloc 0) .--> 1))) & s +* (I +* (Initialize ((intloc 0) .--> 1))) = s +* I & (s +* (I +* (Initialize ((intloc 0) .--> 1)))) +* (Directed I) = s +* (Directed I) )
proof end;

theorem Th21: :: SCM_HALT:21
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitClosed Program of SCM+FSA st p +* I halts_on s +* I & Directed I c= s & Directed I c= p & Initialize ((intloc 0) .--> 1) c= s holds
IC (Comput (p,s,((LifeSpan ((p +* I),(s +* I))) + 1))) = card I
proof end;

theorem Th22: :: SCM_HALT:22
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitClosed Program of SCM+FSA st p +* I halts_on s +* I & Directed I c= s & Directed I c= p & Initialize ((intloc 0) .--> 1) c= s holds
DataPart (Comput (p,s,(LifeSpan ((p +* I),(s +* I))))) = DataPart (Comput (p,s,((LifeSpan ((p +* I),(s +* I))) + 1)))
proof end;

theorem Th23: :: SCM_HALT:23
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting Program of SCM+FSA st Initialized I c= s & I c= p holds
for k being Element of NAT st k <= LifeSpan (p,s) holds
CurInstr ((p +* (Directed I)),(Comput ((p +* (Directed I)),(s +* (Directed I)),k))) <> halt SCM+FSA
proof end;

theorem Th24: :: SCM_HALT:24
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitClosed Program of SCM+FSA st p +* I halts_on s +* (Initialized I) holds
for J being Program of SCM+FSA
for k being Element of NAT st k <= LifeSpan ((p +* I),(s +* (Initialized I))) holds
Comput ((p +* I),(s +* (Initialized I)),k), Comput ((p +* (I ';' J)),(s +* (Initialized (I ';' J))),k) equal_outside NAT
proof end;

theorem Th25: :: SCM_HALT:25
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for J being InitHalting Program of SCM+FSA
for s being State of SCM+FSA st Initialized (I ';' J) c= s & I ';' J c= p holds
( IC (Comput (p,s,((LifeSpan ((p +* I),(s +* I))) + 1))) = card I & DataPart (Comput (p,s,((LifeSpan ((p +* I),(s +* I))) + 1))) = DataPart ((Comput ((p +* I),(s +* I),(LifeSpan ((p +* I),(s +* I))))) +* (Initialized J)) & Reloc (J,(card I)) c= Comput (p,s,((LifeSpan ((p +* I),(s +* I))) + 1)) & (Comput (p,s,((LifeSpan ((p +* I),(s +* I))) + 1))) . (intloc 0) = 1 & p halts_on s & LifeSpan (p,s) = ((LifeSpan ((p +* I),(s +* I))) + 1) + (LifeSpan (((p +* I) +* J),((Result ((p +* I),(s +* I))) +* (Initialized J)))) & ( J is keeping_0 implies (Result (p,s)) . (intloc 0) = 1 ) )
proof end;

registration
let I be InitHalting keepInt0_1 Program of SCM+FSA;
let J be InitHalting Program of SCM+FSA;
cluster I ';' J -> InitHalting ;
coherence
I ';' J is InitHalting
proof end;
end;

theorem Th26: :: SCM_HALT:26
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being keepInt0_1 Program of SCM+FSA st p +* I halts_on s +* I holds
for J being InitClosed Program of SCM+FSA st Initialized (I ';' J) c= s & I ';' J c= p holds
for k being Element of NAT holds (Comput (((p +* I) +* J),((Result ((p +* I),(s +* I))) +* (Initialized J)),k)) +* (Start-At (((IC (Comput (((p +* I) +* J),((Result ((p +* I),(s +* I))) +* (Initialized J)),k))) + (card I)),SCM+FSA)), Comput ((p +* (I ';' J)),(s +* (I ';' J)),(((LifeSpan ((p +* I),(s +* I))) + 1) + k)) equal_outside NAT
proof end;

theorem Th27: :: SCM_HALT:27
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being keepInt0_1 Program of SCM+FSA st not p +* I halts_on s +* (Initialized I) holds
for J being Program of SCM+FSA
for k being Element of NAT holds Comput ((p +* I),(s +* (Initialized I)),k), Comput ((p +* (I ';' J)),(s +* (Initialized (I ';' J))),k) equal_outside NAT
proof end;

theorem Th28: :: SCM_HALT:28
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for J being InitHalting Program of SCM+FSA holds LifeSpan ((p +* (I ';' J)),(s +* (Initialized (I ';' J)))) = ((LifeSpan ((p +* I),(s +* (Initialized I)))) + 1) + (LifeSpan (((p +* I) +* J),((Result ((p +* I),(s +* (Initialized I)))) +* (Initialized J))))
proof end;

theorem Th29: :: SCM_HALT:29
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for J being InitHalting Program of SCM+FSA holds IExec ((I ';' J),p,s) = (IExec (J,p,(IExec (I,p,s)))) +* (Start-At (((IC (IExec (J,p,(IExec (I,p,s))))) + (card I)),SCM+FSA))
proof end;

registration
let i be parahalting Instruction of SCM+FSA;
cluster Macro i -> InitHalting ;
coherence
Macro i is InitHalting
;
end;

registration
let i be parahalting Instruction of SCM+FSA;
let J be parahalting Program of SCM+FSA;
cluster i ';' J -> InitHalting ;
coherence
i ';' J is InitHalting
;
end;

registration
let i be parahalting keeping_0 Instruction of SCM+FSA;
let J be InitHalting Program of SCM+FSA;
cluster i ';' J -> InitHalting ;
coherence
i ';' J is InitHalting
;
end;

registration
let I, J be keepInt0_1 Program of SCM+FSA;
cluster I ';' J -> keepInt0_1 ;
coherence
I ';' J is keepInt0_1
proof end;
end;

registration
let j be parahalting keeping_0 Instruction of SCM+FSA;
let I be InitHalting keepInt0_1 Program of SCM+FSA;
cluster I ';' j -> InitHalting keepInt0_1 ;
coherence
( I ';' j is InitHalting & I ';' j is keepInt0_1 )
;
end;

registration
let i be parahalting keeping_0 Instruction of SCM+FSA;
let J be InitHalting keepInt0_1 Program of SCM+FSA;
cluster i ';' J -> InitHalting keepInt0_1 ;
coherence
( i ';' J is InitHalting & i ';' J is keepInt0_1 )
;
end;

registration
let j be parahalting Instruction of SCM+FSA;
let I be parahalting Program of SCM+FSA;
cluster I ';' j -> InitHalting ;
coherence
I ';' j is InitHalting
;
end;

registration
let i, j be parahalting Instruction of SCM+FSA;
cluster i ';' j -> InitHalting ;
coherence
i ';' j is InitHalting
;
end;

theorem Th30: :: SCM_HALT:30
for s being State of SCM+FSA
for a being Int-Location
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for J being InitHalting Program of SCM+FSA holds (IExec ((I ';' J),p,s)) . a = (IExec (J,p,(IExec (I,p,s)))) . a
proof end;

theorem Th31: :: SCM_HALT:31
for s being State of SCM+FSA
for f being FinSeq-Location
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for J being InitHalting Program of SCM+FSA holds (IExec ((I ';' J),p,s)) . f = (IExec (J,p,(IExec (I,p,s)))) . f
proof end;

theorem Th32: :: SCM_HALT:32
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for s being State of SCM+FSA holds DataPart (Initialized (IExec (I,p,s))) = DataPart (IExec (I,p,s))
proof end;

theorem Th33: :: SCM_HALT:33
for s being State of SCM+FSA
for a being Int-Location
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for j being parahalting Instruction of SCM+FSA holds (IExec ((I ';' j),p,s)) . a = (Exec (j,(IExec (I,p,s)))) . a
proof end;

theorem Th34: :: SCM_HALT:34
for s being State of SCM+FSA
for f being FinSeq-Location
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being InitHalting keepInt0_1 Program of SCM+FSA
for j being parahalting Instruction of SCM+FSA holds (IExec ((I ';' j),p,s)) . f = (Exec (j,(IExec (I,p,s)))) . f
proof end;

definition
let I be Program of SCM+FSA;
let s be State of SCM+FSA;
let p be the Instructions of SCM+FSA -valued ManySortedSet of NAT ;
pred I is_closed_onInit s,p means :Def4: :: SCM_HALT:def 4
for k being Element of NAT holds IC (Comput ((p +* I),(s +* (Initialized I)),k)) in dom I;
pred I is_halting_onInit s,p means :Def5: :: SCM_HALT:def 5
p +* I halts_on s +* (Initialized I);
end;

:: deftheorem Def4 defines is_closed_onInit SCM_HALT:def 4 :
for I being Program of SCM+FSA
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT holds
( I is_closed_onInit s,p iff for k being Element of NAT holds IC (Comput ((p +* I),(s +* (Initialized I)),k)) in dom I );

:: deftheorem Def5 defines is_halting_onInit SCM_HALT:def 5 :
for I being Program of SCM+FSA
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT holds
( I is_halting_onInit s,p iff p +* I halts_on s +* (Initialized I) );

theorem Th35: :: SCM_HALT:35
for I being Program of SCM+FSA holds
( I is InitClosed iff for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT holds I is_closed_onInit s,p )
proof end;

theorem Th36: :: SCM_HALT:36
for I being Program of SCM+FSA holds
( I is InitHalting iff for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT holds I is_halting_onInit s,p )
proof end;

theorem Th37: :: SCM_HALT:37
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being Program of SCM+FSA
for a being Int-Location st not I destroys a & I is_closed_onInit s,p & Initialized I c= s & I c= p holds
for k being Element of NAT holds (Comput (p,s,k)) . a = s . a
proof end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() good V162() InitHalting set ;
existence
ex b1 being Program of SCM+FSA st
( b1 is InitHalting & b1 is good )
proof end;
end;

registration
cluster non empty Relation-like NAT -defined the carrier of SCM+FSA -defined the Instructions of SCM+FSA -valued Function-like the Object-Kind of SCM+FSA -compatible V27() V63() good InitClosed -> keepInt0_1 set ;
correctness
coherence
for b1 being Program of SCM+FSA st b1 is InitClosed & b1 is good holds
b1 is keepInt0_1
;
proof end;
end;

registration
cluster Stop SCM+FSA -> good InitHalting ;
coherence
( Stop SCM+FSA is InitHalting & Stop SCM+FSA is good )
;
end;

theorem :: SCM_HALT:38
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for i being parahalting keeping_0 Instruction of SCM+FSA
for J being InitHalting 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 end;

theorem :: SCM_HALT:39
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for i being parahalting keeping_0 Instruction of SCM+FSA
for J being InitHalting 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 end;

theorem Th40: :: SCM_HALT:40
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being Program of SCM+FSA holds
( I is_closed_onInit s,p iff I is_closed_on Initialized s,p )
proof end;

theorem Th41: :: SCM_HALT:41
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being Program of SCM+FSA holds
( I is_halting_onInit s,p iff I is_halting_on Initialized s,p )
proof end;

theorem :: SCM_HALT:42
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being Program of SCM+FSA
for s being State of SCM+FSA holds IExec (I,p,s) = IExec (I,p,(Initialized s))
proof end;

theorem Th43: :: SCM_HALT:43
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & I is_halting_onInit s,p holds
( if=0 (a,I,J) is_closed_onInit s,p & if=0 (a,I,J) is_halting_onInit s,p )
proof end;

theorem Th44: :: SCM_HALT:44
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & I is_halting_onInit 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 end;

theorem Th45: :: SCM_HALT:45
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & J is_halting_onInit s,p holds
( if=0 (a,I,J) is_closed_onInit s,p & if=0 (a,I,J) is_halting_onInit s,p )
proof end;

theorem Th46: :: SCM_HALT:46
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & J is_halting_onInit 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 end;

theorem Th47: :: SCM_HALT:47
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I, J being InitHalting Program of SCM+FSA
for a being read-write Int-Location holds
( if=0 (a,I,J) is InitHalting & ( 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 end;

theorem :: SCM_HALT:48
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I, J being InitHalting 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 end;

theorem Th49: :: SCM_HALT:49
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & I is_halting_onInit s,p holds
( if>0 (a,I,J) is_closed_onInit s,p & if>0 (a,I,J) is_halting_onInit s,p )
proof end;

theorem Th50: :: SCM_HALT:50
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & I is_halting_onInit 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 end;

theorem Th51: :: SCM_HALT:51
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & J is_halting_onInit s,p holds
( if>0 (a,I,J) is_closed_onInit s,p & if>0 (a,I,J) is_halting_onInit s,p )
proof end;

theorem Th52: :: SCM_HALT:52
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & J is_halting_onInit 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 end;

theorem Th53: :: SCM_HALT:53
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I, J being InitHalting Program of SCM+FSA
for a being read-write Int-Location holds
( if>0 (a,I,J) is InitHalting & ( 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 end;

theorem :: SCM_HALT:54
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I, J being InitHalting 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 end;

theorem Th55: :: SCM_HALT:55
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & I is_halting_onInit 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 end;

theorem Th56: :: SCM_HALT:56
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & J is_halting_onInit 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 end;

theorem Th57: :: SCM_HALT:57
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
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_onInit s,p & J is_halting_onInit 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 end;

theorem Th58: :: SCM_HALT:58
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I, J being InitHalting Program of SCM+FSA
for a being read-write Int-Location holds
( if<0 (a,I,J) is InitHalting & ( 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 end;

registration
let I, J be InitHalting Program of SCM+FSA;
let a be read-write Int-Location ;
cluster if=0 (a,I,J) -> InitHalting ;
correctness
coherence
if=0 (a,I,J) is InitHalting
;
by Th47;
cluster if>0 (a,I,J) -> InitHalting ;
correctness
coherence
if>0 (a,I,J) is InitHalting
;
by Th53;
cluster if<0 (a,I,J) -> InitHalting ;
correctness
coherence
if<0 (a,I,J) is InitHalting
;
by Th58;
end;

theorem Th59: :: SCM_HALT:59
for I being Program of SCM+FSA holds
( I is InitHalting iff for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT holds I is_halting_on Initialized s,p )
proof end;

theorem Th60: :: SCM_HALT:60
for I being Program of SCM+FSA holds
( I is InitClosed iff for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT holds I is_closed_on Initialized s,p )
proof end;

theorem Th61: :: SCM_HALT:61
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being InitHalting Program of SCM+FSA
for a being read-write Int-Location holds (IExec (I,p,s)) . a = (Comput ((p +* I),((Initialized s) +* (Initialize I)),(LifeSpan ((p +* I),((Initialized s) +* (Initialize I)))))) . a
proof end;

theorem Th62: :: SCM_HALT:62
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being InitHalting Program of SCM+FSA
for a being Int-Location
for k being Element of NAT st not I destroys a holds
(IExec (I,p,s)) . a = (Comput ((p +* I),((Initialized s) +* (Initialize I)),k)) . a
proof end;

set A = NAT ;

set D = Data-Locations SCM+FSA;

theorem Th63: :: SCM_HALT:63
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being InitHalting Program of SCM+FSA
for a being Int-Location st not I destroys a holds
(IExec (I,p,s)) . a = (Initialized s) . a
proof end;

theorem Th64: :: SCM_HALT:64
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being InitHalting keepInt0_1 Program of SCM+FSA
for a being read-write Int-Location st not I destroys a holds
(Comput ((p +* (I ';' (SubFrom (a,(intloc 0))))),((Initialized s) +* (Initialize (I ';' (SubFrom (a,(intloc 0)))))),(LifeSpan ((p +* (I ';' (SubFrom (a,(intloc 0))))),((Initialized s) +* (Initialize (I ';' (SubFrom (a,(intloc 0)))))))))) . a = (s . a) - 1
proof end;

theorem Th65: :: SCM_HALT:65
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being InitClosed Program of SCM+FSA st Initialized I c= s & I c= p & p halts_on s holds
for m being Element of NAT st m <= LifeSpan (p,s) holds
Comput (p,s,m), Comput ((p +* (loop I)),(s +* (loop I)),m) equal_outside NAT
proof end;

theorem :: SCM_HALT:66
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being InitHalting Program of SCM+FSA st Initialized I c= s & I c= p holds
for k being Element of NAT st k <= LifeSpan (p,s) holds
CurInstr ((p +* (loop I)),(Comput ((p +* (loop I)),(s +* (loop I)),k))) <> halt SCM+FSA
proof end;

theorem :: SCM_HALT:67
for I being Program of SCM+FSA
for s being State of SCM+FSA holds I c= s +* (Initialized I)
proof end;

theorem Th68: :: SCM_HALT:68
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being Program of SCM+FSA st I is_closed_onInit s,p & I is_halting_onInit s,p holds
for m being Element of NAT st m <= LifeSpan ((p +* I),(s +* (Initialized I))) holds
Comput ((p +* I),(s +* (Initialized I)),m), Comput ((p +* (loop I)),(s +* (Initialized (loop I))),m) equal_outside NAT
proof end;

theorem Th69: :: SCM_HALT:69
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being Program of SCM+FSA st I is_closed_onInit s,p & I is_halting_onInit s,p holds
for m being Element of NAT st m < LifeSpan ((p +* I),(s +* (Initialized I))) holds
CurInstr ((p +* I),(Comput ((p +* I),(s +* (Initialized I)),m))) = CurInstr ((p +* (loop I)),(Comput ((p +* (loop I)),(s +* (Initialized (loop I))),m)))
proof end;

theorem :: SCM_HALT:70
for l being Element of NAT holds not l in dom (Initialize ((intloc 0) .--> 1))
proof end;

theorem Th71: :: SCM_HALT:71
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being Program of SCM+FSA st I is_closed_onInit s,p & I is_halting_onInit s,p holds
( CurInstr ((p +* (loop I)),(Comput ((p +* (loop I)),(s +* (Initialized (loop I))),(LifeSpan ((p +* I),(s +* (Initialized I))))))) = goto 0 & ( for m being Element of NAT st m <= LifeSpan ((p +* I),(s +* (Initialized I))) holds
CurInstr ((p +* (loop I)),(Comput ((p +* (loop I)),(s +* (Initialized (loop I))),m))) <> halt SCM+FSA ) )
proof end;

theorem :: SCM_HALT:72
canceled;

theorem Th73: :: SCM_HALT:73
for s being State of SCM+FSA
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st not I destroys a & s . (intloc 0) = 1 & s . a > 0 holds
loop (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0)))))) is_pseudo-closed_on s,p
proof end;

theorem :: SCM_HALT:74
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st not I destroys a & s . a > 0 holds
Initialized (loop (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0))))))) is_pseudo-closed_on s,p
proof end;

theorem :: SCM_HALT:75
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st not I destroys a & s . (intloc 0) = 1 holds
( Times (a,I) is_closed_on s,p & Times (a,I) is_halting_on s,p )
proof end;

theorem :: SCM_HALT:76
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st not I destroys a holds
Initialized (Times (a,I)) is halting
proof end;

registration
let a be read-write Int-Location ;
let I be good Program of SCM+FSA;
cluster Times (a,I) -> good ;
coherence
Times (a,I) is good
;
end;

theorem Th77: :: SCM_HALT:77
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st not I destroys a & s . a > 0 holds
ex s2 being State of SCM+FSA ex p2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT ex k being Element of NAT st
( s2 = s +* (Initialized (loop (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0)))))))) & p2 = p +* (loop (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0))))))) & k = (LifeSpan ((p +* (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0))))))),(s +* (Initialized (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0)))))))))) + 1 & (Comput (p2,s2,k)) . a = (s . a) - 1 & (Comput (p2,s2,k)) . (intloc 0) = 1 & ( for b being read-write Int-Location st b <> a holds
(Comput (p2,s2,k)) . b = (IExec (I,p,s)) . b ) & ( for f being FinSeq-Location holds (Comput (p2,s2,k)) . f = (IExec (I,p,s)) . f ) & IC (Comput (p2,s2,k)) = 0 & ( for n being Element of NAT st n <= k holds
IC (Comput (p2,s2,n)) in dom (loop (if=0 (a,(Goto 2),(I ';' (SubFrom (a,(intloc 0))))))) ) )
proof end;

theorem Th78: :: SCM_HALT:78
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st s . (intloc 0) = 1 & s . a <= 0 holds
DataPart (IExec ((Times (a,I)),p,s)) = DataPart s
proof end;

theorem Th79: :: SCM_HALT:79
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for a being read-write Int-Location st not I destroys a & s . a > 0 holds
( (IExec ((I ';' (SubFrom (a,(intloc 0)))),p,s)) . a = (s . a) - 1 & DataPart (IExec ((Times (a,I)),p,s)) = DataPart (IExec ((Times (a,I)),p,(IExec ((I ';' (SubFrom (a,(intloc 0)))),p,s)))) )
proof end;

theorem :: SCM_HALT:80
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for f being FinSeq-Location
for a being read-write Int-Location st s . a <= 0 holds
(IExec ((Times (a,I)),p,s)) . f = s . f
proof end;

theorem :: SCM_HALT:81
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for b being Int-Location
for a being read-write Int-Location st s . a <= 0 holds
(IExec ((Times (a,I)),p,s)) . b = (Initialized s) . b
proof end;

theorem :: SCM_HALT:82
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for f being FinSeq-Location
for a being read-write Int-Location st not I destroys a & s . a > 0 holds
(IExec ((Times (a,I)),p,s)) . f = (IExec ((Times (a,I)),p,(IExec ((I ';' (SubFrom (a,(intloc 0)))),p,s)))) . f
proof end;

theorem :: SCM_HALT:83
for p being the Instructions of SCM+FSA -valued ManySortedSet of NAT
for s being State of SCM+FSA
for I being good InitHalting Program of SCM+FSA
for b being Int-Location
for a being read-write Int-Location st not I destroys a & s . a > 0 holds
(IExec ((Times (a,I)),p,s)) . b = (IExec ((Times (a,I)),p,(IExec ((I ';' (SubFrom (a,(intloc 0)))),p,s)))) . b
proof end;

definition
let i be Instruction of SCM+FSA;
redefine attr i is good means :Def6: :: SCM_HALT:def 6
not i destroys intloc 0;
compatibility
( i is good iff not i destroys intloc 0 )
proof end;
end;

:: deftheorem Def6 defines good SCM_HALT:def 6 :
for i being Instruction of SCM+FSA holds
( i is good iff not i destroys intloc 0 );

registration
cluster good parahalting with_explicit_jumps IC-relocable Exec-preserving Element of the Instructions of SCM+FSA;
existence
ex b1 being Instruction of SCM+FSA st
( b1 is parahalting & b1 is good )
proof end;
end;

registration
let n be Element of NAT ;
cluster intloc (n + 1) -> read-write ;
coherence
not intloc (n + 1) is read-only
proof end;
end;