:: Computation in { \bf SCM_FSA }
:: by Andrzej Trybulec and Yatsuka Nakamura
::
:: Received February 7, 1996
:: Copyright (c) 1996-2011 Association of Mizar Users


begin

theorem :: SCMFSA_3:1
not IC in Int-Locations
proof end;

theorem :: SCMFSA_3:2
not IC in FinSeq-Locations
proof end;

begin

definition
let s be State of SCM+FSA;
let p be preProgram of SCM+FSA;
:: original: +*
redefine func s +* p -> State of SCM+FSA;
coherence
s +* p is State of SCM+FSA
;
end;

theorem :: SCMFSA_3:3
canceled;

theorem :: SCMFSA_3:4
canceled;

theorem :: SCMFSA_3:5
canceled;

theorem :: SCMFSA_3:6
canceled;

theorem :: SCMFSA_3:7
canceled;

theorem :: SCMFSA_3:8
canceled;

theorem :: SCMFSA_3:9
canceled;

theorem :: SCMFSA_3:10
canceled;

theorem :: SCMFSA_3:11
for s being State of SCM+FSA
for iloc being Element of NAT
for a being Int-Location holds s . a = (s +* (Start-At (iloc,SCM+FSA))) . a
proof end;

theorem :: SCMFSA_3:12
for s being State of SCM+FSA
for iloc being Element of NAT
for a being FinSeq-Location holds s . a = (s +* (Start-At (iloc,SCM+FSA))) . a
proof end;

begin

definition
let la be Int-Location ;
let a be Integer;
:: original: .-->
redefine func la .--> a -> FinPartState of SCM+FSA;
coherence
la .--> a is FinPartState of SCM+FSA
proof end;
end;

registration
cluster SCM+FSA -> IC-recognized ;
coherence
SCM+FSA is IC-recognized
proof end;
end;

theorem :: SCMFSA_3:13
canceled;

theorem :: SCMFSA_3:14
canceled;

theorem :: SCMFSA_3:15
canceled;

theorem :: SCMFSA_3:16
canceled;

registration
cluster SCM+FSA -> CurIns-recognized ;
coherence
SCM+FSA is CurIns-recognized
proof end;
end;

theorem :: SCMFSA_3:17
canceled;

theorem :: SCMFSA_3:18
canceled;

Lm1: for da being Int-Location
for p being PartState of SCM+FSA st da in dom p holds
da in dom (NPP p)
proof end;

Lm2: for fa being FinSeq-Location
for p being PartState of SCM+FSA st fa in dom p holds
fa in dom (NPP p)
proof end;

theorem :: SCMFSA_3:19
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location st CurInstr (P1,(Comput (P1,s1,i))) = da := db & da in dom p holds
(Comput (P1,s1,i)) . db = (Comput (P2,s2,i)) . db
proof end;

theorem :: SCMFSA_3:20
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location st CurInstr (P1,(Comput (P1,s1,i))) = AddTo (da,db) & da in dom p holds
((Comput (P1,s1,i)) . da) + ((Comput (P1,s1,i)) . db) = ((Comput (P2,s2,i)) . da) + ((Comput (P2,s2,i)) . db)
proof end;

theorem :: SCMFSA_3:21
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location st CurInstr (P1,(Comput (P1,s1,i))) = SubFrom (da,db) & da in dom p holds
((Comput (P1,s1,i)) . da) - ((Comput (P1,s1,i)) . db) = ((Comput (P2,s2,i)) . da) - ((Comput (P2,s2,i)) . db)
proof end;

theorem :: SCMFSA_3:22
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location st CurInstr (P1,(Comput (P1,s1,i))) = MultBy (da,db) & da in dom p holds
((Comput (P1,s1,i)) . da) * ((Comput (P1,s1,i)) . db) = ((Comput (P2,s2,i)) . da) * ((Comput (P2,s2,i)) . db)
proof end;

theorem :: SCMFSA_3:23
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location st CurInstr (P1,(Comput (P1,s1,i))) = Divide (da,db) & da in dom p & da <> db holds
((Comput (P1,s1,i)) . da) div ((Comput (P1,s1,i)) . db) = ((Comput (P2,s2,i)) . da) div ((Comput (P2,s2,i)) . db)
proof end;

theorem :: SCMFSA_3:24
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location st CurInstr (P1,(Comput (P1,s1,i))) = Divide (da,db) & db in dom p holds
((Comput (P1,s1,i)) . da) mod ((Comput (P1,s1,i)) . db) = ((Comput (P2,s2,i)) . da) mod ((Comput (P2,s2,i)) . db)
proof end;

theorem :: SCMFSA_3:25
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da being Int-Location
for loc being Element of NAT st CurInstr (P1,(Comput (P1,s1,i))) = da =0_goto loc & loc <> succ (IC (Comput (P1,s1,i))) holds
( (Comput (P1,s1,i)) . da = 0 iff (Comput (P2,s2,i)) . da = 0 )
proof end;

theorem :: SCMFSA_3:26
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da being Int-Location
for loc being Element of NAT st CurInstr (P1,(Comput (P1,s1,i))) = da >0_goto loc & loc <> succ (IC (Comput (P1,s1,i))) holds
( (Comput (P1,s1,i)) . da > 0 iff (Comput (P2,s2,i)) . da > 0 )
proof end;

theorem :: SCMFSA_3:27
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location
for f being FinSeq-Location st CurInstr (P1,(Comput (P1,s1,i))) = da := (f,db) & da in dom p holds
for k1, k2 being Element of NAT st k1 = abs ((Comput (P1,s1,i)) . db) & k2 = abs ((Comput (P2,s2,i)) . db) holds
((Comput (P1,s1,i)) . f) /. k1 = ((Comput (P2,s2,i)) . f) /. k2
proof end;

theorem :: SCMFSA_3:28
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da, db being Int-Location
for f being FinSeq-Location st CurInstr (P1,(Comput (P1,s1,i))) = (f,db) := da & f in dom p holds
for k1, k2 being Element of NAT st k1 = abs ((Comput (P1,s1,i)) . db) & k2 = abs ((Comput (P2,s2,i)) . db) holds
((Comput (P1,s1,i)) . f) +* (k1,((Comput (P1,s1,i)) . da)) = ((Comput (P2,s2,i)) . f) +* (k2,((Comput (P2,s2,i)) . da))
proof end;

theorem :: SCMFSA_3:29
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da being Int-Location
for f being FinSeq-Location st CurInstr (P1,(Comput (P1,s1,i))) = da :=len f & da in dom p holds
len ((Comput (P1,s1,i)) . f) = len ((Comput (P2,s2,i)) . f)
proof end;

theorem :: SCMFSA_3:30
for p being non NAT -defined autonomic FinPartState of
for s1, s2 being State of SCM+FSA st NPP p c= s1 & NPP p c= s2 holds
for P1, P2 being the Instructions of SCM+FSA -valued ManySortedSet of NAT st ProgramPart p c= P1 & ProgramPart p c= P2 holds
for i being Element of NAT
for da being Int-Location
for f being FinSeq-Location st CurInstr (P1,(Comput (P1,s1,i))) = f :=<0,...,0> da & f in dom p holds
for k1, k2 being Element of NAT st k1 = abs ((Comput (P1,s1,i)) . da) & k2 = abs ((Comput (P2,s2,i)) . da) holds
k1 |-> 0 = k2 |-> 0
proof end;