:: AMISTD_1 semantic presentation

theorem Th1: :: AMISTD_1:1
for r being real number holds max {r} = r
proof end;

theorem Th2: :: AMISTD_1:2
for n being Element of NAT holds max {n} = n by ;

theorem Th3: :: AMISTD_1:3
for D being non empty set
for f being trivial FinSequence of D holds
( f is empty or ex x being Element of D st f = <*x*> )
proof end;

registration
cluster -> with_non-empty_elements set ;
coherence
for b1 being Relation holds b1 is with_non-empty_elements
proof end;
end;

registration
let A be finite set ;
let B be set ;
cluster K225(a1,a2) -> finite ;
coherence
A --> B is finite
proof end;
end;

registration
let f1 be non empty FinSequence;
let f2 be FinSequence;
cluster a1 ^' a2 -> non empty with_non-empty_elements ;
coherence
not f1 ^' f2 is empty
proof end;
end;

theorem Th4: :: AMISTD_1:4
canceled;

theorem Th5: :: AMISTD_1:5
canceled;

theorem Th6: :: AMISTD_1:6
canceled;

theorem Th7: :: AMISTD_1:7
canceled;

theorem Th8: :: AMISTD_1:8
canceled;

theorem Th9: :: AMISTD_1:9
canceled;

theorem Th10: :: AMISTD_1:10
canceled;

theorem Th11: :: AMISTD_1:11
for N being with_non-empty_elements set
for S being non empty non void definite AMI-Struct of N
for I being Element of the Instructions of S
for s being State of S holds s +* (the Instruction-Locations of S --> I) is State of S
proof end;

registration
let N be set ;
let S be AMI-Struct of N;
cluster empty -> with_non-empty_elements programmed FinPartState of a2;
coherence
for b1 being FinPartState of S st b1 is empty holds
b1 is programmed
proof end;
end;

registration
let N be set ;
let S be AMI-Struct of N;
cluster empty with_non-empty_elements programmed FinPartState of a2;
existence
ex b1 being FinPartState of S st b1 is empty
proof end;
end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
cluster non empty trivial with_non-empty_elements programmed FinPartState of a2;
existence
ex b1 being FinPartState of S st
( not b1 is empty & b1 is trivial & b1 is programmed )
proof end;
end;

registration
let N be with_non-empty_elements set ;
let S be non void AMI-Struct of N;
let i be Element of the Instructions of S;
let s be State of S;
cluster (the Execution of a2 . a3) . a4 -> Relation-like Function-like with_non-empty_elements ;
coherence
( (the Execution of S . i) . s is Function-like & (the Execution of S . i) . s is Relation-like )
proof end;
end;

Lemma60: for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated steady-programmed definite AMI-Struct of N
for il being Instruction-Location of S
for I being Element of the Instructions of S
for f being FinPartState of S st f = il .--> I holds
f is autonomic
proof end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated steady-programmed definite AMI-Struct of N;
cluster non empty trivial with_non-empty_elements autonomic programmed FinPartState of a2;
existence
ex b1 being FinPartState of S st
( not b1 is empty & b1 is trivial & b1 is autonomic & b1 is programmed )
proof end;
end;

theorem Th12: :: AMISTD_1:12
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated steady-programmed definite AMI-Struct of N
for il being Instruction-Location of S
for I being Element of the Instructions of S holds il .--> I is autonomic by ;

theorem Th13: :: AMISTD_1:13
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated steady-programmed definite AMI-Struct of N holds S is programmable
proof end;

registration
let N be with_non-empty_elements set ;
cluster non empty non void IC-Ins-separated steady-programmed definite -> non empty non void IC-Ins-separated definite programmable AMI-Struct of a1;
coherence
for b1 being non empty non void IC-Ins-separated definite AMI-Struct of N st b1 is steady-programmed holds
b1 is programmable
by ;
end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void AMI-Struct of N;
let T be InsType of S;
canceled;
canceled;
attr a3 is jump-only means :: AMISTD_1:def 3
for s being State of S
for o being Object of S
for I being Instruction of S st InsCode I = T & o <> IC S holds
(Exec I,s) . o = s . o;
end;

:: deftheorem Def1 AMISTD_1:def 1 :
canceled;

:: deftheorem Def2 AMISTD_1:def 2 :
canceled;

:: deftheorem Def3 defines jump-only AMISTD_1:def 3 :
for N being with_non-empty_elements set
for S being non empty non void AMI-Struct of N
for T being InsType of S holds
( T is jump-only iff for s being State of S
for o being Object of S
for I being Instruction of S st InsCode I = T & o <> IC S holds
(Exec I,s) . o = s . o );

definition
let N be with_non-empty_elements set ;
let S be non empty non void AMI-Struct of N;
let I be Instruction of S;
attr a3 is jump-only means :: AMISTD_1:def 4
InsCode I is jump-only;
end;

:: deftheorem Def4 defines jump-only AMISTD_1:def 4 :
for N being with_non-empty_elements set
for S being non empty non void AMI-Struct of N
for I being Instruction of S holds
( I is jump-only iff InsCode I is jump-only );

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
let l be Instruction-Location of S;
let i be Element of the Instructions of S;
func NIC c4,c3 -> Subset of the Instruction-Locations of a2 equals :: AMISTD_1:def 5
{ (IC (Following s)) where s is State of S : ( IC s = l & s . l = i ) } ;
coherence
{ (IC (Following s)) where s is State of S : ( IC s = l & s . l = i ) } is Subset of the Instruction-Locations of S
proof end;
end;

:: deftheorem Def5 defines NIC AMISTD_1:def 5 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for l being Instruction-Location of S
for i being Element of the Instructions of S holds NIC i,l = { (IC (Following s)) where s is State of S : ( IC s = l & s . l = i ) } ;

E77: now
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite realistic AMI-Struct of N;
let i be Element of the Instructions of S;
let l be Instruction-Location of S;
let s be State of S;
let f be FinPartState of S;
assume E36: f = (IC S),l --> l,i ;
set t = s +* f;
E37: IC S <> l by AMI_1:48;
E39: dom f = {(IC S),l} by , FUNCT_4:65;
then E40: IC S in dom f by TARSKI:def 2;
E41: IC (s +* f) = f . (IC S) by , FUNCT_4:14
.= l by , , FUNCT_4:66 ;
l in dom f by , TARSKI:def 2;
then (s +* f) . l = f . l by FUNCT_4:14
.= i by , , FUNCT_4:66 ;
hence IC (Following (s +* f)) in NIC i,l by ;
end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite realistic AMI-Struct of N;
let i be Element of the Instructions of S;
let l be Instruction-Location of S;
cluster NIC a3,a4 -> non empty ;
coherence
not NIC i,l is empty
proof end;
end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
let i be Element of the Instructions of S;
func JUMP c3 -> Subset of the Instruction-Locations of a2 equals :: AMISTD_1:def 6
meet { (NIC i,l) where l is Instruction-Location of S : verum } ;
coherence
meet { (NIC i,l) where l is Instruction-Location of S : verum } is Subset of the Instruction-Locations of S
proof end;
end;

:: deftheorem Def6 defines JUMP AMISTD_1:def 6 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for i being Element of the Instructions of S holds JUMP i = meet { (NIC i,l) where l is Instruction-Location of S : verum } ;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
let l be Instruction-Location of S;
func SUCC c3 -> Subset of the Instruction-Locations of a2 equals :: AMISTD_1:def 7
union { ((NIC i,l) \ (JUMP i)) where i is Element of the Instructions of S : verum } ;
coherence
union { ((NIC i,l) \ (JUMP i)) where i is Element of the Instructions of S : verum } is Subset of the Instruction-Locations of S
proof end;
end;

:: deftheorem Def7 defines SUCC AMISTD_1:def 7 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for l being Instruction-Location of S holds SUCC l = union { ((NIC i,l) \ (JUMP i)) where i is Element of the Instructions of S : verum } ;

theorem Th14: :: AMISTD_1:14
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for i being Element of the Instructions of S st not the Instruction-Locations of S is trivial & ( for l being Instruction-Location of S holds NIC i,l = {l} ) holds
JUMP i is empty
proof end;

theorem Th15: :: AMISTD_1:15
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite realistic AMI-Struct of N
for il being Instruction-Location of S
for i being Instruction of S st i is halting holds
NIC i,il = {il}
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
let l1 be Instruction-Location of S;
let l2 be Instruction-Location of S;
pred c3 <= c4 means :Def8: :: AMISTD_1:def 8
ex f being non empty FinSequence of the Instruction-Locations of S st
( f /. 1 = l1 & f /. (len f) = l2 & ( for n being Element of NAT st 1 <= n & n < len f holds
f /. (n + 1) in SUCC (f /. n) ) );
reflexivity
for l1 being Instruction-Location of S ex f being non empty FinSequence of the Instruction-Locations of S st
( f /. 1 = l1 & f /. (len f) = l1 & ( for n being Element of NAT st 1 <= n & n < len f holds
f /. (n + 1) in SUCC (f /. n) ) )
proof end;
end;

:: deftheorem Def8 defines <= AMISTD_1:def 8 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for l1, l2 being Instruction-Location of S holds
( l1 <= l2 iff ex f being non empty FinSequence of the Instruction-Locations of S st
( f /. 1 = l1 & f /. (len f) = l2 & ( for n being Element of NAT st 1 <= n & n < len f holds
f /. (n + 1) in SUCC (f /. n) ) ) );

theorem Th16: :: AMISTD_1:16
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for l1, l2, l3 being Instruction-Location of S st l1 <= l2 & l2 <= l3 holds
l1 <= l3
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
attr a2 is InsLoc-antisymmetric means :: AMISTD_1:def 9
for l1, l2 being Instruction-Location of S st l1 <= l2 & l2 <= l1 holds
l1 = l2;
end;

:: deftheorem Def9 defines InsLoc-antisymmetric AMISTD_1:def 9 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N holds
( S is InsLoc-antisymmetric iff for l1, l2 being Instruction-Location of S st l1 <= l2 & l2 <= l1 holds
l1 = l2 );

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
attr a2 is standard means :Def10: :: AMISTD_1:def 10
ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) );
end;

:: deftheorem Def10 defines standard AMISTD_1:def 10 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N holds
( S is standard iff ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) ) );

theorem Th17: :: AMISTD_1:17
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for f1, f2 being Function of NAT ,the Instruction-Locations of S st f1 is bijective & ( for m, n being Element of NAT holds
( m <= n iff f1 . m <= f1 . n ) ) & f2 is bijective & ( for m, n being Element of NAT holds
( m <= n iff f2 . m <= f2 . n ) ) holds
f1 = f2
proof end;

theorem Th18: :: AMISTD_1:18
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for f being Function of NAT ,the Instruction-Locations of S st f is bijective holds
( ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) iff for k being Element of NAT holds
( f . (k + 1) in SUCC (f . k) & ( for j being Element of NAT st f . j in SUCC (f . k) holds
k <= j ) ) )
proof end;

theorem Th19: :: AMISTD_1:19
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N holds
( S is standard iff ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for k being Element of NAT holds
( f . (k + 1) in SUCC (f . k) & ( for j being Element of NAT st f . j in SUCC (f . k) holds
k <= j ) ) ) ) )
proof end;

Lemma164: for a, b being set holds dom ((NAT --> a) +* (NAT .--> b)) = NAT \/ {NAT }
proof end;

set III = {[1,0],[0,0]};

definition
let N be with_non-empty_elements set ;
func STC c1 -> strict AMI-Struct of a1 means :Def11: :: AMISTD_1:def 11
( the carrier of it = NAT \/ {NAT } & the Instruction-Counter of it = NAT & the Instruction-Locations of it = NAT & the Instruction-Codes of it = {0,1} & the Instructions of it = {[0,0],[1,0]} & the Object-Kind of it = (NAT --> {[1,0],[0,0]}) +* (NAT .--> NAT ) & ex f being Function of product the Object-Kind of it, product the Object-Kind of it st
( ( for s being Element of product the Object-Kind of it holds f . s = s +* (NAT .--> (succ (s . NAT ))) ) & the Execution of it = ([1,0] .--> f) +* ([0,0] .--> (id (product the Object-Kind of it))) ) );
existence
ex b1 being strict AMI-Struct of N st
( the carrier of b1 = NAT \/ {NAT } & the Instruction-Counter of b1 = NAT & the Instruction-Locations of b1 = NAT & the Instruction-Codes of b1 = {0,1} & the Instructions of b1 = {[0,0],[1,0]} & the Object-Kind of b1 = (NAT --> {[1,0],[0,0]}) +* (NAT .--> NAT ) & ex f being Function of product the Object-Kind of b1, product the Object-Kind of b1 st
( ( for s being Element of product the Object-Kind of b1 holds f . s = s +* (NAT .--> (succ (s . NAT ))) ) & the Execution of b1 = ([1,0] .--> f) +* ([0,0] .--> (id (product the Object-Kind of b1))) ) )
proof end;
uniqueness
for b1, b2 being strict AMI-Struct of N st the carrier of b1 = NAT \/ {NAT } & the Instruction-Counter of b1 = NAT & the Instruction-Locations of b1 = NAT & the Instruction-Codes of b1 = {0,1} & the Instructions of b1 = {[0,0],[1,0]} & the Object-Kind of b1 = (NAT --> {[1,0],[0,0]}) +* (NAT .--> NAT ) & ex f being Function of product the Object-Kind of b1, product the Object-Kind of b1 st
( ( for s being Element of product the Object-Kind of b1 holds f . s = s +* (NAT .--> (succ (s . NAT ))) ) & the Execution of b1 = ([1,0] .--> f) +* ([0,0] .--> (id (product the Object-Kind of b1))) ) & the carrier of b2 = NAT \/ {NAT } & the Instruction-Counter of b2 = NAT & the Instruction-Locations of b2 = NAT & the Instruction-Codes of b2 = {0,1} & the Instructions of b2 = {[0,0],[1,0]} & the Object-Kind of b2 = (NAT --> {[1,0],[0,0]}) +* (NAT .--> NAT ) & ex f being Function of product the Object-Kind of b2, product the Object-Kind of b2 st
( ( for s being Element of product the Object-Kind of b2 holds f . s = s +* (NAT .--> (succ (s . NAT ))) ) & the Execution of b2 = ([1,0] .--> f) +* ([0,0] .--> (id (product the Object-Kind of b2))) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def11 defines STC AMISTD_1:def 11 :
for N being with_non-empty_elements set
for b2 being strict AMI-Struct of N holds
( b2 = STC N iff ( the carrier of b2 = NAT \/ {NAT } & the Instruction-Counter of b2 = NAT & the Instruction-Locations of b2 = NAT & the Instruction-Codes of b2 = {0,1} & the Instructions of b2 = {[0,0],[1,0]} & the Object-Kind of b2 = (NAT --> {[1,0],[0,0]}) +* (NAT .--> NAT ) & ex f being Function of product the Object-Kind of b2, product the Object-Kind of b2 st
( ( for s being Element of product the Object-Kind of b2 holds f . s = s +* (NAT .--> (succ (s . NAT ))) ) & the Execution of b2 = ([1,0] .--> f) +* ([0,0] .--> (id (product the Object-Kind of b2))) ) ) );

registration
let N be with_non-empty_elements set ;
cluster the Instruction-Locations of (STC a1) -> infinite ;
coherence
not the Instruction-Locations of (STC N) is finite
by , CARD_4:15;
end;

registration
let N be with_non-empty_elements set ;
cluster STC a1 -> non empty strict non void ;
coherence
( not STC N is empty & not STC N is void )
proof end;
end;

registration
let N be with_non-empty_elements set ;
cluster STC a1 -> non empty strict non void IC-Ins-separated data-oriented steady-programmed definite realistic programmable ;
coherence
( STC N is IC-Ins-separated & STC N is definite & STC N is realistic & STC N is steady-programmed & STC N is data-oriented )
proof end;
end;

Lemma176: for N being with_non-empty_elements set
for i being Instruction of (STC N)
for s being State of (STC N) st InsCode i = 1 holds
(Exec i,s) . (IC (STC N)) = succ (IC s)
proof end;

theorem Th20: :: AMISTD_1:20
for N being with_non-empty_elements set
for i being Instruction of (STC N) st InsCode i = 0 holds
i is halting
proof end;

theorem Th21: :: AMISTD_1:21
for N being with_non-empty_elements set
for i being Instruction of (STC N) st InsCode i = 1 holds
not i is halting
proof end;

theorem Th22: :: AMISTD_1:22
for N being with_non-empty_elements set
for i being Element of the Instructions of (STC N) holds
( InsCode i = 1 or InsCode i = 0 )
proof end;

theorem Th23: :: AMISTD_1:23
for N being with_non-empty_elements set
for i being Instruction of (STC N) holds i is jump-only
proof end;

Lemma180: for z being natural number
for N being with_non-empty_elements set
for l being Instruction-Location of (STC N)
for i being Element of the Instructions of (STC N) st l = z & InsCode i = 1 holds
NIC i,l = {(z + 1)}
proof end;

Lemma184: for N being with_non-empty_elements set
for i being Element of the Instructions of (STC N) holds JUMP i is empty
proof end;

theorem Th24: :: AMISTD_1:24
for z being natural number
for N being with_non-empty_elements set
for l being Instruction-Location of (STC N) st l = z holds
SUCC l = {z,(z + 1)}
proof end;

registration
let N be with_non-empty_elements set ;
cluster STC a1 -> non empty strict non void IC-Ins-separated data-oriented steady-programmed definite realistic programmable standard ;
coherence
STC N is standard
proof end;
end;

registration
let N be with_non-empty_elements set ;
cluster STC a1 -> non empty strict non void halting IC-Ins-separated data-oriented steady-programmed definite realistic programmable standard ;
coherence
STC N is halting
proof end;
end;

registration
let N be with_non-empty_elements set ;
cluster non empty non void halting IC-Ins-separated steady-programmed definite realistic programmable standard AMI-Struct of a1;
existence
ex b1 being non empty non void IC-Ins-separated definite AMI-Struct of N st
( b1 is standard & b1 is halting & b1 is realistic & b1 is steady-programmed & b1 is programmable )
proof end;
end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let k be natural number ;
func il. c2,c3 -> Instruction-Location of a2 means :Def12: :: AMISTD_1:def 12
ex f being Function of NAT ,the Instruction-Locations of N st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) & it = f . S );
existence
ex b1 being Instruction-Location of S ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) & b1 = f . k )
proof end;
uniqueness
for b1, b2 being Instruction-Location of S st ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) & b1 = f . k ) & ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) & b2 = f . k ) holds
b1 = b2
by ;
end;

:: deftheorem Def12 defines il. AMISTD_1:def 12 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for k being natural number
for b4 being Instruction-Location of S holds
( b4 = il. S,k iff ex f being Function of NAT ,the Instruction-Locations of S st
( f is bijective & ( for m, n being Element of NAT holds
( m <= n iff f . m <= f . n ) ) & b4 = f . k ) );

theorem Th25: :: AMISTD_1:25
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for k1, k2 being natural number st il. T,k1 = il. T,k2 holds
k1 = k2
proof end;

theorem Th26: :: AMISTD_1:26
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for l being Instruction-Location of T ex k being natural number st l = il. T,k
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let l be Instruction-Location of S;
func locnum c3 -> natural number means :Def13: :: AMISTD_1:def 13
il. N,it = S;
existence
ex b1 being natural number st il. S,b1 = l
by ;
uniqueness
for b1, b2 being natural number st il. S,b1 = l & il. S,b2 = l holds
b1 = b2
by ;
end;

:: deftheorem Def13 defines locnum AMISTD_1:def 13 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for l being Instruction-Location of S
for b4 being natural number holds
( b4 = locnum l iff il. S,b4 = l );

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let l be Instruction-Location of S;
redefine func locnum as locnum c3 -> Element of NAT ;
coherence
locnum l is Element of NAT
by ORDINAL1:def 13;
end;

theorem Th27: :: AMISTD_1:27
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for l1, l2 being Instruction-Location of T st locnum l1 = locnum l2 holds
l1 = l2
proof end;

theorem Th28: :: AMISTD_1:28
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for k1, k2 being natural number holds
( il. T,k1 <= il. T,k2 iff k1 <= k2 )
proof end;

theorem Th29: :: AMISTD_1:29
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for l1, l2 being Instruction-Location of T holds
( locnum l1 <= locnum l2 iff l1 <= l2 )
proof end;

theorem Th30: :: AMISTD_1:30
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N st S is standard holds
S is InsLoc-antisymmetric
proof end;

registration
let N be with_non-empty_elements set ;
cluster non empty non void IC-Ins-separated definite standard -> non empty non void IC-Ins-separated definite InsLoc-antisymmetric AMI-Struct of a1;
coherence
for b1 being non empty non void IC-Ins-separated definite AMI-Struct of N st b1 is standard holds
b1 is InsLoc-antisymmetric
by ;
end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let f be Instruction-Location of S;
let k be natural number ;
func c3 + c4 -> Instruction-Location of a2 equals :: AMISTD_1:def 14
il. N,((locnum S) + f);
coherence
il. S,((locnum f) + k) is Instruction-Location of S
;
end;

:: deftheorem Def14 defines + AMISTD_1:def 14 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f being Instruction-Location of S
for k being natural number holds f + k = il. S,((locnum f) + k);

theorem Th31: :: AMISTD_1:31
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f being Instruction-Location of T holds f + 0 = f by ;

theorem Th32: :: AMISTD_1:32
for z being natural number
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f, g being Instruction-Location of T st f + z = g + z holds
f = g
proof end;

theorem Th33: :: AMISTD_1:33
for z being natural number
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f being Instruction-Location of T holds (locnum f) + z = locnum (f + z) by ;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let f be Instruction-Location of S;
func NextLoc c3 -> Instruction-Location of a2 equals :: AMISTD_1:def 15
S + 1;
coherence
f + 1 is Instruction-Location of S
;
end;

:: deftheorem Def15 defines NextLoc AMISTD_1:def 15 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f being Instruction-Location of S holds NextLoc f = f + 1;

theorem Th34: :: AMISTD_1:34
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f being Instruction-Location of T holds NextLoc f = il. T,((locnum f) + 1) ;

theorem Th35: :: AMISTD_1:35
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f being Instruction-Location of T holds f <> NextLoc f
proof end;

theorem Th36: :: AMISTD_1:36
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for f, g being Instruction-Location of T st NextLoc f = NextLoc g holds
f = g
proof end;

theorem Th37: :: AMISTD_1:37
for z being natural number
for N being with_non-empty_elements set holds il. (STC N),z = z
proof end;

theorem Th38: :: AMISTD_1:38
for N being with_non-empty_elements set
for i being Instruction of (STC N)
for s being State of (STC N) st InsCode i = 1 holds
(Exec i,s) . (IC (STC N)) = NextLoc (IC s)
proof end;

theorem Th39: :: AMISTD_1:39
for N being with_non-empty_elements set
for l being Instruction-Location of (STC N)
for i being Element of the Instructions of (STC N) st InsCode i = 1 holds
NIC i,l = {(NextLoc l)}
proof end;

theorem Th40: :: AMISTD_1:40
for N being with_non-empty_elements set
for l being Instruction-Location of (STC N) holds SUCC l = {l,(NextLoc l)}
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let i be Instruction of S;
attr a3 is sequential means :: AMISTD_1:def 16
for s being State of N holds (Exec S,s) . (IC N) = NextLoc (IC s);
end;

:: deftheorem Def16 defines sequential AMISTD_1:def 16 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for i being Instruction of S holds
( i is sequential iff for s being State of S holds (Exec i,s) . (IC S) = NextLoc (IC s) );

theorem Th41: :: AMISTD_1:41
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite realistic standard AMI-Struct of N
for il being Instruction-Location of S
for i being Instruction of S st i is sequential holds
NIC i,il = {(NextLoc il)}
proof end;

theorem Th42: :: AMISTD_1:42
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite realistic standard AMI-Struct of N
for i being Instruction of S st i is sequential holds
not i is halting
proof end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite realistic standard AMI-Struct of N;
cluster sequential -> non halting Element of the Instructions of a2;
coherence
for b1 being Instruction of S st b1 is sequential holds
not b1 is halting
by Th3;
cluster halting -> non sequential Element of the Instructions of a2;
coherence
for b1 being Instruction of S st b1 is halting holds
not b1 is sequential
by Th3;
end;

theorem Th43: :: AMISTD_1:43
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for i being Instruction of T st not JUMP i is empty holds
not i is sequential
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
let F be FinPartState of S;
attr a3 is closed means :Def17: :: AMISTD_1:def 17
for l being Instruction-Location of N st l in dom S holds
NIC (pi S,l),l c= dom S;
attr a3 is really-closed means :: AMISTD_1:def 18
for s being State of N st S c= s & IC s in dom S holds
for k being Element of NAT holds IC ((Computation s) . k) in dom S;
end;

:: deftheorem Def17 defines closed AMISTD_1:def 17 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for F being FinPartState of S holds
( F is closed iff for l being Instruction-Location of S st l in dom F holds
NIC (pi F,l),l c= dom F );

:: deftheorem Def18 defines really-closed AMISTD_1:def 18 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for F being FinPartState of S holds
( F is really-closed iff for s being State of S st F c= s & IC s in dom F holds
for k being Element of NAT holds IC ((Computation s) . k) in dom F );

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let F be FinPartState of S;
attr a3 is para-closed means :: AMISTD_1:def 19
for s being State of N st S c= s & IC s = il. N,0 holds
for k being Element of NAT holds IC ((Computation s) . k) in dom S;
end;

:: deftheorem Def19 defines para-closed AMISTD_1:def 19 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being FinPartState of S holds
( F is para-closed iff for s being State of S st F c= s & IC s = il. S,0 holds
for k being Element of NAT holds IC ((Computation s) . k) in dom F );

theorem Th44: :: AMISTD_1:44
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated steady-programmed definite standard AMI-Struct of N
for F being FinPartState of S st F is really-closed & il. S,0 in dom F holds
F is para-closed
proof end;

theorem Th45: :: AMISTD_1:45
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated steady-programmed definite AMI-Struct of N
for F being FinPartState of S st F is closed holds
F is really-closed
proof end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated steady-programmed definite AMI-Struct of N;
cluster closed -> with_non-empty_elements really-closed FinPartState of a2;
coherence
for b1 being FinPartState of S st b1 is closed holds
b1 is really-closed
by ;
end;

theorem Th46: :: AMISTD_1:46
for N being with_non-empty_elements set
for S being non empty non void halting IC-Ins-separated definite realistic standard AMI-Struct of N holds (il. S,0) .--> (halt S) is closed
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
let F be FinPartState of S;
attr a3 is lower means :Def20: :: AMISTD_1:def 20
for l being Instruction-Location of N st l in dom S holds
for m being Instruction-Location of N st m <= l holds
m in dom S;
end;

:: deftheorem Def20 defines lower AMISTD_1:def 20 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for F being FinPartState of S holds
( F is lower iff for l being Instruction-Location of S st l in dom F holds
for m being Instruction-Location of S st m <= l holds
m in dom F );

theorem Th47: :: AMISTD_1:47
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite AMI-Struct of N
for F being empty FinPartState of S holds F is lower
proof end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite AMI-Struct of N;
cluster empty -> with_non-empty_elements lower FinPartState of a2;
coherence
for b1 being FinPartState of S st b1 is empty holds
b1 is lower
by ;
end;

theorem Th48: :: AMISTD_1:48
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for i being Element of the Instructions of T holds (il. T,0) .--> i is lower
proof end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
cluster non empty trivial with_non-empty_elements programmed lower FinPartState of a2;
existence
ex b1 being FinPartState of S st
( b1 is lower & not b1 is empty & b1 is trivial & b1 is programmed )
proof end;
end;

theorem Th49: :: AMISTD_1:49
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed lower FinPartState of T holds il. T,0 in dom F
proof end;

theorem Th50: :: AMISTD_1:50
for z being natural number
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for P being programmed lower FinPartState of T holds
( z < card P iff il. T,z in dom P )
proof end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated definite standard AMI-Struct of N;
let F be non empty programmed FinPartState of S;
func LastLoc c3 -> Instruction-Location of a2 means :Def21: :: AMISTD_1:def 21
ex M being non empty finite natural-membered set st
( M = { (locnum l) where l is Element of the Instruction-Locations of N : l in dom S } & it = il. N,(max M) );
existence
ex b1 being Instruction-Location of S ex M being non empty finite natural-membered set st
( M = { (locnum l) where l is Element of the Instruction-Locations of S : l in dom F } & b1 = il. S,(max M) )
proof end;
uniqueness
for b1, b2 being Instruction-Location of S st ex M being non empty finite natural-membered set st
( M = { (locnum l) where l is Element of the Instruction-Locations of S : l in dom F } & b1 = il. S,(max M) ) & ex M being non empty finite natural-membered set st
( M = { (locnum l) where l is Element of the Instruction-Locations of S : l in dom F } & b2 = il. S,(max M) ) holds
b1 = b2
;
end;

:: deftheorem Def21 defines LastLoc AMISTD_1:def 21 :
for N being with_non-empty_elements set
for S being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed FinPartState of S
for b4 being Instruction-Location of S holds
( b4 = LastLoc F iff ex M being non empty finite natural-membered set st
( M = { (locnum l) where l is Element of the Instruction-Locations of S : l in dom F } & b4 = il. S,(max M) ) );

theorem Th51: :: AMISTD_1:51
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed FinPartState of T holds LastLoc F in dom F
proof end;

theorem Th52: :: AMISTD_1:52
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F, G being non empty programmed FinPartState of T st F c= G holds
LastLoc F <= LastLoc G
proof end;

theorem Th53: :: AMISTD_1:53
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed FinPartState of T
for l being Instruction-Location of T st l in dom F holds
l <= LastLoc F
proof end;

theorem Th54: :: AMISTD_1:54
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed lower FinPartState of T
for G being non empty programmed FinPartState of T st F c= G & LastLoc F = LastLoc G holds
F = G
proof end;

theorem Th55: :: AMISTD_1:55
for N being with_non-empty_elements set
for T being non empty non void IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed lower FinPartState of T holds LastLoc F = il. T,((card F) -' 1)
proof end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void IC-Ins-separated steady-programmed definite standard AMI-Struct of N;
cluster non empty programmed really-closed lower -> with_non-empty_elements para-closed FinPartState of a2;
coherence
for b1 being FinPartState of S st b1 is really-closed & b1 is lower & not b1 is empty & b1 is programmed holds
b1 is para-closed
proof end;
end;

E220: now
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite standard AMI-Struct of N;
set F = (il. S,0) .--> (halt S);
E36: dom ((il. S,0) .--> (halt S)) = {(il. S,0)} by CQC_LANG:5;
then E37: card (dom ((il. S,0) .--> (halt S))) = 1 by CARD_1:79;
(il. S,0) .--> (halt S) is lower FinPartState of S by ;
then E39: LastLoc ((il. S,0) .--> (halt S)) = il. S,((card ((il. S,0) .--> (halt S))) -' 1) by
.= il. S,((card (dom ((il. S,0) .--> (halt S)))) -' 1) by PRE_CIRC:21
.= il. S,0 by , BINARITH:51 ;
hence ((il. S,0) .--> (halt S)) . (LastLoc ((il. S,0) .--> (halt S))) = halt S by CQC_LANG:6;
let l be Instruction-Location of S;
assume ((il. S,0) .--> (halt S)) . l = halt S ;
assume l in dom ((il. S,0) .--> (halt S)) ;
hence l = LastLoc ((il. S,0) .--> (halt S)) by , , TARSKI:def 1;
end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite standard AMI-Struct of N;
let F be non empty programmed FinPartState of S;
attr a3 is halt-ending means :: AMISTD_1:def 22
S . (LastLoc S) = halt N;
attr a3 is unique-halt means :: AMISTD_1:def 23
for f being Instruction-Location of N st S . f = halt N & f in dom S holds
f = LastLoc S;
end;

:: deftheorem Def22 defines halt-ending AMISTD_1:def 22 :
for N being with_non-empty_elements set
for S being non empty non void halting IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed FinPartState of S holds
( F is halt-ending iff F . (LastLoc F) = halt S );

:: deftheorem Def23 defines unique-halt AMISTD_1:def 23 :
for N being with_non-empty_elements set
for S being non empty non void halting IC-Ins-separated definite standard AMI-Struct of N
for F being non empty programmed FinPartState of S holds
( F is unique-halt iff for f being Instruction-Location of S st F . f = halt S & f in dom F holds
f = LastLoc F );

registration
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite standard AMI-Struct of N;
cluster non empty trivial with_non-empty_elements programmed lower halt-ending unique-halt FinPartState of a2;
existence
ex b1 being non empty programmed lower FinPartState of S st
( b1 is halt-ending & b1 is unique-halt & b1 is trivial )
proof end;
end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite realistic standard AMI-Struct of N;
cluster non empty trivial with_non-empty_elements programmed closed lower FinPartState of a2;
existence
ex b1 being FinPartState of S st
( b1 is trivial & b1 is closed & b1 is lower & not b1 is empty & b1 is programmed )
proof end;
end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite realistic standard AMI-Struct of N;
cluster non empty trivial with_non-empty_elements programmed closed lower halt-ending unique-halt FinPartState of a2;
existence
ex b1 being non empty programmed lower FinPartState of S st
( b1 is halt-ending & b1 is unique-halt & b1 is trivial & b1 is closed )
proof end;
end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated steady-programmed definite realistic standard AMI-Struct of N;
cluster non empty trivial with_non-empty_elements autonomic programmed closed really-closed para-closed lower halt-ending unique-halt FinPartState of a2;
existence
ex b1 being non empty programmed lower FinPartState of S st
( b1 is halt-ending & b1 is unique-halt & b1 is autonomic & b1 is trivial & b1 is closed )
proof end;
end;

definition
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite standard AMI-Struct of N;
mode pre-Macro of a2 is non empty programmed lower halt-ending unique-halt FinPartState of a2;
end;

registration
let N be with_non-empty_elements set ;
let S be non empty non void halting IC-Ins-separated definite realistic standard AMI-Struct of N;
cluster closed FinPartState of a2;
existence
ex b1 being pre-Macro of S st b1 is closed
proof end;
end;