:: Standard Ordering of Instruction Locations
:: by Andrzej Trybulec , Piotr Rudnicki and Artur Korni{\l}owicz
::
:: Received April 14, 2000
:: Copyright (c) 2000 Association of Mizar Users


begin

theorem :: AMISTD_1:1
canceled;

theorem :: AMISTD_1:2
canceled;

theorem :: AMISTD_1:3
canceled;

theorem :: AMISTD_1:4
canceled;

theorem :: AMISTD_1:5
canceled;

theorem :: AMISTD_1:6
canceled;

theorem :: AMISTD_1:7
canceled;

theorem :: AMISTD_1:8
canceled;

theorem :: AMISTD_1:9
canceled;

theorem :: AMISTD_1:10
canceled;

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

registration
let N be set ;
let S be AMI-Struct of N;
cluster Relation-like the carrier of S -defined Function-like the Object-Kind of S -compatible empty finite -> NAT -defined set ;
coherence
for b1 being FinPartState of S st b1 is empty holds
b1 is NAT -defined
proof end;
end;

registration
let N be set ;
let S be AMI-Struct of N;
cluster Relation-like the carrier of S -defined Function-like the Object-Kind of S -compatible empty finite countable set ;
existence
ex b1 being FinPartState of S st b1 is empty
proof end;
end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable set ;
existence
ex b1 being FinPartState of S st
( not b1 is empty & b1 is trivial & b1 is NAT -defined )
proof end;
end;

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

Lm1: for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated steady-programmed definite AMI-Struct of N
for il being Element of NAT
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 non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated steady-programmed definite AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable autonomic set ;
existence
ex b1 being FinPartState of S st
( not b1 is empty & b1 is trivial & b1 is autonomic & b1 is NAT -defined )
proof end;
end;

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

definition
let N be with_non-empty_elements set ;
let S be non empty stored-program standard-ins AMI-Struct of N;
let T be InsType of S;
canceled;
canceled;
attr T 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 AMISTD_1:def 1 :
canceled;

:: deftheorem AMISTD_1:def 2 :
canceled;

:: deftheorem defines jump-only AMISTD_1:def 3 :
for N being with_non-empty_elements set
for S being non empty stored-program standard-ins 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 stored-program standard-ins AMI-Struct of N;
let I be Instruction of S;
attr I is jump-only means :: AMISTD_1:def 4
InsCode I is jump-only ;
end;

:: deftheorem defines jump-only AMISTD_1:def 4 :
for N being with_non-empty_elements set
for S being non empty stored-program standard-ins 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 non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite AMI-Struct of N;
let l be Nat;
let i be Element of the Instructions of S;
func NIC i,l -> Subset of NAT equals :: AMISTD_1:def 5
{ (IC (Following (ProgramPart ss),ss)) where ss is Element of product the Object-Kind of S : ( IC ss = l & (ProgramPart ss) /. l = i ) } ;
coherence
{ (IC (Following (ProgramPart ss),ss)) where ss is Element of product the Object-Kind of S : ( IC ss = l & (ProgramPart ss) /. l = i ) } is Subset of NAT
proof end;
end;

:: deftheorem defines NIC AMISTD_1:def 5 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N
for l being Nat
for i being Element of the Instructions of S holds NIC i,l = { (IC (Following (ProgramPart ss),ss)) where ss is Element of product the Object-Kind of S : ( IC ss = l & (ProgramPart ss) /. l = i ) } ;

Lm2: now
let N be non empty with_non-empty_elements set ; :: thesis: for S being non empty stored-program IC-Ins-separated steady-programmed definite realistic AMI-Struct of N
for i being Element of the Instructions of S
for l being Element of NAT
for s being State of S
for f being FinPartState of S st f = (IC S),l --> l,i holds
IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l

let S be non empty stored-program IC-Ins-separated steady-programmed definite realistic AMI-Struct of N; :: thesis: for i being Element of the Instructions of S
for l being Element of NAT
for s being State of S
for f being FinPartState of S st f = (IC S),l --> l,i holds
IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l

let i be Element of the Instructions of S; :: thesis: for l being Element of NAT
for s being State of S
for f being FinPartState of S st f = (IC S),l --> l,i holds
IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l

let l be Element of NAT ; :: thesis: for s being State of S
for f being FinPartState of S st f = (IC S),l --> l,i holds
IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l

let s be State of S; :: thesis: for f being FinPartState of S st f = (IC S),l --> l,i holds
IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l

let f be FinPartState of S; :: thesis: ( f = (IC S),l --> l,i implies IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l )
assume A1: f = (IC S),l --> l,i ; :: thesis: IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l
reconsider t = s +* f as Element of product the Object-Kind of S by PBOOLE:155;
A2: dom f = {(IC S),l} by A1, FUNCT_4:65;
then X: l in dom f by TARSKI:def 2;
A3: (ProgramPart t) /. l = t . l by AMI_1:150
.= f . l by FUNCT_4:14, X
.= i by A1, FUNCT_4:66 ;
IC S in dom f by A2, TARSKI:def 2;
then IC t = f . (IC S) by FUNCT_4:14
.= l by A1, AMI_1:48, FUNCT_4:66 ;
hence IC (Following (ProgramPart (s +* f)),(s +* f)) in NIC i,l by A3; :: thesis: verum
end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated steady-programmed definite realistic AMI-Struct of N;
let i be Element of the Instructions of S;
let l be Element of NAT ;
cluster NIC i,l -> non empty ;
coherence
not NIC i,l is empty
proof end;
end;

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

:: deftheorem defines JUMP AMISTD_1:def 6 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program 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 Element of NAT : verum } ;

definition
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite AMI-Struct of N;
let l be Nat;
func SUCC l,S -> Subset of NAT 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 NAT
proof end;
end;

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

theorem :: AMISTD_1:13
canceled;

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

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

begin

definition
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite AMI-Struct of N;
let l1, l2 be Nat;
pred l1 <= l2,S means :Def8: :: AMISTD_1:def 8
ex f being non empty FinSequence of NAT 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),S ) );
end;

:: deftheorem Def8 defines <= AMISTD_1:def 8 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N
for l1, l2 being Nat holds
( l1 <= l2,S iff ex f being non empty FinSequence of NAT 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),S ) ) );

theorem :: AMISTD_1:16
for l3 being Element of NAT
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N
for l1, l2 being Nat st l1 <= l2,S & l2 <= l3,S holds
l1 <= l3,S
proof end;

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

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

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

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

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

LemRefle: for k being Element of NAT
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N holds k <= k,S
proof end;

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

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

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

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

begin

definition
let N be with_non-empty_elements set ;
func STC N -> strict AMI-Struct of N means :Def11: :: AMISTD_1:def 11
( the carrier of it = NAT \/ {NAT } & the Instruction-Counter of it = NAT & 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 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 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 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 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 STC N -> infinite strict ;
coherence
not STC N is finite
proof end;
end;

registration
let N be with_non-empty_elements set ;
cluster STC N -> non empty strict stored-program standard-ins ;
coherence
( not STC N is empty & STC N is stored-program & STC N is standard-ins )
proof end;
end;

registration
let N be with_non-empty_elements set ;
cluster STC N -> strict IC-Ins-separated steady-programmed definite realistic ;
coherence
( STC N is IC-Ins-separated & STC N is definite & STC N is realistic & STC N is steady-programmed )
proof end;
end;

Lm4: for N being non empty 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 non empty with_non-empty_elements set
for i being Instruction of (STC N) st InsCode i = 0 holds
i is halting
proof end;

theorem :: AMISTD_1:21
for N being non empty 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 non empty 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 :: AMISTD_1:23
for N being non empty with_non-empty_elements set
for i being Instruction of (STC N) holds i is jump-only
proof end;

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

Lm6: for N being non empty 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 non empty with_non-empty_elements set
for l being Element of NAT st l = z holds
SUCC l,(STC N) = {z,(z + 1)}
proof end;

registration
let N be non empty with_non-empty_elements set ;
cluster STC N -> strict standard ;
coherence
STC N is standard
proof end;
end;

registration
let N be non empty with_non-empty_elements set ;
cluster STC N -> strict halting ;
coherence
STC N is halting
proof end;
end;

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

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

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

theorem Th25: :: AMISTD_1:25
for N being non empty with_non-empty_elements set
for T being non empty stored-program 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 non empty with_non-empty_elements set
for T being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for l being Nat ex k being natural number st l = il. T,k
proof end;

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

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

definition
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite standard AMI-Struct of N;
let l be Nat;
:: original: locnum
redefine func locnum l,S -> Element of NAT ;
coherence
locnum l,S is Element of NAT
by ORDINAL1:def 13;
end;

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

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

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

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

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

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

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

theorem :: AMISTD_1:31
for N being non empty with_non-empty_elements set
for T being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for f being Element of NAT holds f + 0 ,T = f by Def13;

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

theorem :: AMISTD_1:33
for z being natural number
for N being non empty with_non-empty_elements set
for T being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for f being Element of NAT holds (locnum f,T) + z = locnum (f + z,T),T by Def13;

definition
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite standard AMI-Struct of N;
let f be Element of NAT ;
func NextLoc f,S -> Element of NAT equals :: AMISTD_1:def 15
f + 1,S;
coherence
f + 1,S is Element of NAT
;
end;

:: deftheorem defines NextLoc AMISTD_1:def 15 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for f being Element of NAT holds NextLoc f,S = f + 1,S;

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

theorem Th35: :: AMISTD_1:35
for N being non empty with_non-empty_elements set
for T being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for f being Element of NAT holds f <> NextLoc f,T
proof end;

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

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

theorem :: AMISTD_1:38
for N being non empty 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),(STC N)
proof end;

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

theorem :: AMISTD_1:40
for N being non empty with_non-empty_elements set
for l being Element of NAT holds SUCC l,(STC N) = {l,(NextLoc l,(STC N))}
proof end;

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

:: deftheorem defines sequential AMISTD_1:def 16 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program 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),S );

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

theorem Th42: :: AMISTD_1:42
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated steady-programmed 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 non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated steady-programmed definite realistic standard AMI-Struct of N;
cluster sequential -> non halting Element of the Instructions of S;
coherence
for b1 being Instruction of S st b1 is sequential holds
not b1 is halting
by Th42;
cluster halting -> non sequential Element of the Instructions of S;
coherence
for b1 being Instruction of S st b1 is halting holds
not b1 is sequential
;
end;

theorem :: AMISTD_1:43
for N being non empty with_non-empty_elements set
for T being non empty stored-program 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;

begin

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

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

:: deftheorem defines really-closed AMISTD_1:def 18 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N
for F being NAT -defined FinPartState of 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 (Comput (ProgramPart s),s,k) in dom F );

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

:: deftheorem defines para-closed AMISTD_1:def 19 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program 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 (Comput (ProgramPart s),s,k) in dom F );

theorem Th44: :: AMISTD_1:44
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated steady-programmed definite standard AMI-Struct of N
for F being NAT -defined FinPartState of 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 non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated steady-programmed definite AMI-Struct of N
for F being NAT -defined FinPartState of st F is closed holds
F is really-closed
proof end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated steady-programmed definite AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible finite closed -> NAT -defined really-closed set ;
coherence
for b1 being NAT -defined FinPartState of st b1 is closed holds
b1 is really-closed
by Th45;
end;

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

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

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

theorem Th47: :: AMISTD_1:47
for N being non empty with_non-empty_elements set
for S being non empty stored-program 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 non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite AMI-Struct of N;
cluster Relation-like the carrier of S -defined Function-like the Object-Kind of S -compatible empty finite -> lower set ;
coherence
for b1 being FinPartState of S st b1 is empty holds
b1 is lower
by Th47;
end;

theorem Th48: :: AMISTD_1:48
for N being non empty with_non-empty_elements set
for T being non empty stored-program 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 non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable lower set ;
existence
ex b1 being FinPartState of S st
( b1 is lower & not b1 is empty & b1 is trivial & b1 is NAT -defined )
proof end;
end;

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

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

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

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

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

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

theorem Th53: :: AMISTD_1:53
for N being non empty with_non-empty_elements set
for T being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for F being NAT -defined non empty FinPartState of
for l being Element of NAT st l in dom F holds
l <= LastLoc F,T
proof end;

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

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

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated steady-programmed definite standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty finite really-closed lower -> NAT -defined para-closed set ;
coherence
for b1 being NAT -defined FinPartState of st b1 is really-closed & b1 is lower & not b1 is empty holds
b1 is para-closed
proof end;
end;

Lm7: now
let N be non empty with_non-empty_elements set ; :: thesis: for S being non empty stored-program halting IC-Ins-separated definite standard AMI-Struct of N holds
( ((il. S,0 ) .--> (halt S)) . (LastLoc ((il. S,0 ) .--> (halt S))) = halt S & ( for l being Element of NAT st ((il. S,0 ) .--> (halt S)) . l = halt S & l in dom ((il. S,0 ) .--> (halt S)) holds
l = LastLoc ((il. S,0 ) .--> (halt S)) ) )

let S be non empty stored-program halting IC-Ins-separated definite standard AMI-Struct of N; :: thesis: ( ((il. S,0 ) .--> (halt S)) . (LastLoc ((il. S,0 ) .--> (halt S))) = halt S & ( for l being Element of NAT st ((il. S,0 ) .--> (halt S)) . l = halt S & l in dom ((il. S,0 ) .--> (halt S)) holds
l = LastLoc ((il. S,0 ) .--> (halt S)) ) )

set F = (il. S,0 ) .--> (halt S);
A1: dom ((il. S,0 ) .--> (halt S)) = {(il. S,0 )} by FUNCOP_1:19;
then A2: card (dom ((il. S,0 ) .--> (halt S))) = 1 by CARD_1:50;
(il. S,0 ) .--> (halt S) is lower FinPartState of S by Th48;
then A3: LastLoc ((il. S,0 ) .--> (halt S)) = il. S,((card ((il. S,0 ) .--> (halt S))) -' 1) by Th55
.= il. S,((card (dom ((il. S,0 ) .--> (halt S)))) -' 1) by CARD_1:104
.= il. S,0 by A2, XREAL_1:234 ;
hence ((il. S,0 ) .--> (halt S)) . (LastLoc ((il. S,0 ) .--> (halt S))) = halt S by FUNCOP_1:87; :: thesis: for l being Element of NAT st ((il. S,0 ) .--> (halt S)) . l = halt S & l in dom ((il. S,0 ) .--> (halt S)) holds
l = LastLoc ((il. S,0 ) .--> (halt S))

let l be Element of NAT ; :: thesis: ( ((il. S,0 ) .--> (halt S)) . l = halt S & l in dom ((il. S,0 ) .--> (halt S)) implies l = LastLoc ((il. S,0 ) .--> (halt S)) )
assume ((il. S,0 ) .--> (halt S)) . l = halt S ; :: thesis: ( l in dom ((il. S,0 ) .--> (halt S)) implies l = LastLoc ((il. S,0 ) .--> (halt S)) )
assume l in dom ((il. S,0 ) .--> (halt S)) ; :: thesis: l = LastLoc ((il. S,0 ) .--> (halt S))
hence l = LastLoc ((il. S,0 ) .--> (halt S)) by A1, A3, TARSKI:def 1; :: thesis: verum
end;

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

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

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

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program halting IC-Ins-separated definite standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable lower halt-ending unique-halt set ;
existence
ex b1 being NAT -defined non empty lower FinPartState of st
( b1 is halt-ending & b1 is unique-halt & b1 is trivial )
proof end;
end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program halting IC-Ins-separated steady-programmed definite realistic standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable closed lower set ;
existence
ex b1 being NAT -defined FinPartState of st
( b1 is trivial & b1 is closed & b1 is lower & not b1 is empty )
proof end;
end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program halting IC-Ins-separated steady-programmed definite realistic standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable closed lower halt-ending unique-halt set ;
existence
ex b1 being NAT -defined non empty lower FinPartState of st
( b1 is halt-ending & b1 is unique-halt & b1 is trivial & b1 is closed )
proof end;
end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program halting IC-Ins-separated steady-programmed definite realistic standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty trivial finite countable autonomic closed lower halt-ending unique-halt set ;
existence
ex b1 being NAT -defined non empty lower FinPartState of 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 non empty with_non-empty_elements set ;
let S be non empty stored-program halting IC-Ins-separated definite standard AMI-Struct of N;
mode pre-Macro of S is NAT -defined non empty lower halt-ending unique-halt FinPartState of ;
end;

registration
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program halting IC-Ins-separated steady-programmed definite realistic standard AMI-Struct of N;
cluster Relation-like NAT -defined the carrier of S -defined Function-like the Object-Kind of S -compatible non empty finite countable closed lower halt-ending unique-halt set ;
existence
ex b1 being pre-Macro of S st b1 is closed
proof end;
end;

theorem :: AMISTD_1:56
for N being non empty with_non-empty_elements set
for i being Instruction of (Trivial-AMI N) holds i is halting
proof end;

theorem :: AMISTD_1:57
for N being non empty with_non-empty_elements set
for i being Element of the Instructions of (Trivial-AMI N) holds InsCode i = 0
proof end;

begin

theorem :: AMISTD_1:58
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N
for i being Instruction of S
for l being Element of NAT holds JUMP i c= NIC i,l
proof end;

theorem :: AMISTD_1:59
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite AMI-Struct of N
for l1, l2 being Element of NAT st SUCC l1,S = NAT holds
l1 <= l2,S
proof end;

definition
let N be non empty with_non-empty_elements set ;
let S be non empty stored-program IC-Ins-separated definite standard AMI-Struct of N;
let loc be Element of NAT ;
let k be natural number ;
func loc -' k,S -> Element of NAT equals :: AMISTD_1:def 24
il. S,((locnum loc,S) -' k);
coherence
il. S,((locnum loc,S) -' k) is Element of NAT
;
end;

:: deftheorem defines -' AMISTD_1:def 24 :
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for loc being Element of NAT
for k being natural number holds loc -' k,S = il. S,((locnum loc,S) -' k);

theorem :: AMISTD_1:60
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for l being Element of NAT holds l -' 0 ,S = l
proof end;

theorem Th61: :: AMISTD_1:61
for k being natural number
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for l being Element of NAT holds (l + k,S) -' k,S = l
proof end;

theorem :: AMISTD_1:62
canceled;

theorem :: AMISTD_1:63
for k being natural number
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for l1, l2 being Element of NAT holds
( Start-At (l1 + k,S),S = Start-At (l2 + k,S),S iff Start-At l1,S = Start-At l2,S )
proof end;

theorem :: AMISTD_1:64
for k being natural number
for N being non empty with_non-empty_elements set
for S being non empty stored-program IC-Ins-separated definite standard AMI-Struct of N
for l1, l2 being Element of NAT st Start-At l1,S = Start-At l2,S holds
Start-At (l1 -' k,S),S = Start-At (l2 -' k,S),S
proof end;