:: SCM_COMP semantic presentation

theorem Th1: :: SCM_COMP:1
for I1, I2 being FinSequence of the Instructions of SCM
for D being FinSequence of INT
for il, ps, ds being Element of NAT
for s being State-consisting of il,ps,ds,I1 ^ I2,D holds
( s is State-consisting of il,ps,ds,I1,D & s is State-consisting of il,ps + (len I1),ds,I2,D )
proof end;

theorem Th2: :: SCM_COMP:2
for I1, I2 being FinSequence of the Instructions of SCM
for il, ps, ds, k, il1 being Element of NAT
for s being State-consisting of il,ps,ds,I1 ^ I2, <*> INT
for u being State of SCM st u = (Computation s) . k & il. il1 = IC u holds
u is State-consisting of il1,ps + (len I1),ds,I2, <*> INT
proof end;

Lemma49: 1 = { n where n is Element of NAT : n < 1 }
by AXIOMS:30;

Lemma51: 5 = { n where n is Element of NAT : n < 5 }
by AXIOMS:30;

definition
func SCM-AE -> non empty strict with_terminals with_nonterminals with_useful_nonterminals binary DTConstrStr means :Def1: :: SCM_COMP:def 1
( Terminals it = SCM-Data-Loc & NonTerminals it = [:1,5:] & ( for x, y, z being Symbol of it holds
( x ==> <*y,z*> iff x in [:1,5:] ) ) );
existence
ex b1 being non empty strict with_terminals with_nonterminals with_useful_nonterminals binary DTConstrStr st
( Terminals b1 = SCM-Data-Loc & NonTerminals b1 = [:1,5:] & ( for x, y, z being Symbol of b1 holds
( x ==> <*y,z*> iff x in [:1,5:] ) ) )
proof end;
uniqueness
for b1, b2 being non empty strict with_terminals with_nonterminals with_useful_nonterminals binary DTConstrStr st Terminals b1 = SCM-Data-Loc & NonTerminals b1 = [:1,5:] & ( for x, y, z being Symbol of b1 holds
( x ==> <*y,z*> iff x in [:1,5:] ) ) & Terminals b2 = SCM-Data-Loc & NonTerminals b2 = [:1,5:] & ( for x, y, z being Symbol of b2 holds
( x ==> <*y,z*> iff x in [:1,5:] ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def1 defines SCM-AE SCM_COMP:def 1 :
for b1 being non empty strict with_terminals with_nonterminals with_useful_nonterminals binary DTConstrStr holds
( b1 = SCM-AE iff ( Terminals b1 = SCM-Data-Loc & NonTerminals b1 = [:1,5:] & ( for x, y, z being Symbol of b1 holds
( x ==> <*y,z*> iff x in [:1,5:] ) ) ) );

definition
mode bin-term is Element of TS SCM-AE ;
end;

Lemma97: NonTerminals SCM-AE = [:1,5:]
by ;

definition
let nt be NonTerminal of SCM-AE ;
let tl be bin-term, tr be bin-term;
redefine func -tree as c1 -tree c2,c3 -> bin-term;
coherence
nt -tree tl,tr is bin-term
proof end;
end;

definition
let t be Terminal of SCM-AE ;
redefine func root-tree as root-tree c1 -> bin-term;
coherence
root-tree t is bin-term
by DTCONSTR:def 4;
end;

definition
let t be Terminal of SCM-AE ;
func @ c1 -> Data-Location equals :: SCM_COMP:def 2
t;
coherence
t is Data-Location
proof end;
end;

:: deftheorem Def2 defines @ SCM_COMP:def 2 :
for t being Terminal of SCM-AE holds @ t = t;

theorem Th3: :: SCM_COMP:3
for nt being NonTerminal of SCM-AE holds
( nt = [0,0] or nt = [0,1] or nt = [0,2] or nt = [0,3] or nt = [0,4] )
proof end;

theorem Th4: :: SCM_COMP:4
( [0,0] is NonTerminal of SCM-AE & [0,1] is NonTerminal of SCM-AE & [0,2] is NonTerminal of SCM-AE & [0,3] is NonTerminal of SCM-AE & [0,4] is NonTerminal of SCM-AE )
proof end;

then reconsider nt0 = [0,0], nt1 = [0,1], nt2 = [0,2], nt3 = [0,3], nt4 = [0,4] as NonTerminal of SCM-AE ;

definition
let t1 be bin-term, t2 be bin-term;
func c1 + c2 -> bin-term equals :: SCM_COMP:def 3
[0,0] -tree nt0,nt1;
coherence
[0,0] -tree t1,t2 is bin-term
proof end;
func c1 - c2 -> bin-term equals :: SCM_COMP:def 4
[0,1] -tree nt0,nt1;
coherence
[0,1] -tree t1,t2 is bin-term
proof end;
func c1 * c2 -> bin-term equals :: SCM_COMP:def 5
[0,2] -tree nt0,nt1;
coherence
[0,2] -tree t1,t2 is bin-term
proof end;
func c1 div c2 -> bin-term equals :: SCM_COMP:def 6
[0,3] -tree nt0,nt1;
coherence
[0,3] -tree t1,t2 is bin-term
proof end;
func c1 mod c2 -> bin-term equals :: SCM_COMP:def 7
[0,4] -tree nt0,nt1;
coherence
[0,4] -tree t1,t2 is bin-term
proof end;
end;

:: deftheorem Def3 defines + SCM_COMP:def 3 :
for t1, t2 being bin-term holds t1 + t2 = [0,0] -tree t1,t2;

:: deftheorem Def4 defines - SCM_COMP:def 4 :
for t1, t2 being bin-term holds t1 - t2 = [0,1] -tree t1,t2;

:: deftheorem Def5 defines * SCM_COMP:def 5 :
for t1, t2 being bin-term holds t1 * t2 = [0,2] -tree t1,t2;

:: deftheorem Def6 defines div SCM_COMP:def 6 :
for t1, t2 being bin-term holds t1 div t2 = [0,3] -tree t1,t2;

:: deftheorem Def7 defines mod SCM_COMP:def 7 :
for t1, t2 being bin-term holds t1 mod t2 = [0,4] -tree t1,t2;

theorem Th5: :: SCM_COMP:5
for term being bin-term holds
( ex t being Terminal of SCM-AE st term = root-tree t or ex tl, tr being bin-term st
( term = tl + tr or term = tl - tr or term = tl * tr or term = tl div tr or term = tl mod tr ) )
proof end;

definition
let o be NonTerminal of SCM-AE ;
let i be Integer, j be Integer;
func c1 -Meaning_on c2,c3 -> Integer equals :Def8: :: SCM_COMP:def 8
nt1 + nt2 if nt0 = [0,0]
nt1 - nt2 if nt0 = [0,1]
nt1 * nt2 if nt0 = [0,2]
nt1 div nt2 if nt0 = [0,3]
nt1 mod nt2 if nt0 = [0,4]
;
coherence
( ( o = [0,0] implies i + j is Integer ) & ( o = [0,1] implies i - j is Integer ) & ( o = [0,2] implies i * j is Integer ) & ( o = [0,3] implies i div j is Integer ) & ( o = [0,4] implies i mod j is Integer ) )
;
consistency
for b1 being Integer holds
( ( o = [0,0] & o = [0,1] implies ( b1 = i + j iff b1 = i - j ) ) & ( o = [0,0] & o = [0,2] implies ( b1 = i + j iff b1 = i * j ) ) & ( o = [0,0] & o = [0,3] implies ( b1 = i + j iff b1 = i div j ) ) & ( o = [0,0] & o = [0,4] implies ( b1 = i + j iff b1 = i mod j ) ) & ( o = [0,1] & o = [0,2] implies ( b1 = i - j iff b1 = i * j ) ) & ( o = [0,1] & o = [0,3] implies ( b1 = i - j iff b1 = i div j ) ) & ( o = [0,1] & o = [0,4] implies ( b1 = i - j iff b1 = i mod j ) ) & ( o = [0,2] & o = [0,3] implies ( b1 = i * j iff b1 = i div j ) ) & ( o = [0,2] & o = [0,4] implies ( b1 = i * j iff b1 = i mod j ) ) & ( o = [0,3] & o = [0,4] implies ( b1 = i div j iff b1 = i mod j ) ) )
proof end;
end;

:: deftheorem Def8 defines -Meaning_on SCM_COMP:def 8 :
for o being NonTerminal of SCM-AE
for i, j being Integer holds
( ( o = [0,0] implies o -Meaning_on i,j = i + j ) & ( o = [0,1] implies o -Meaning_on i,j = i - j ) & ( o = [0,2] implies o -Meaning_on i,j = i * j ) & ( o = [0,3] implies o -Meaning_on i,j = i div j ) & ( o = [0,4] implies o -Meaning_on i,j = i mod j ) );

definition
let s be State of SCM ;
let t be Terminal of SCM-AE ;
redefine func . as c1 . c2 -> Integer;
coherence
s . t is Integer
proof end;
end;

definition
let D be non empty set ;
let f be Function of INT ,D;
let x be Integer;
redefine func . as c2 . c3 -> Element of a1;
coherence
f . x is Element of D
proof end;
end;

set i2i = id INT ;

definition
let s be State of SCM ;
let term be bin-term;
deffunc H1( NonTerminal of SCM-AE , set , set , Integer, Integer) -> Element of INT = (id INT ) . (a1 -Meaning_on a4,a5);
deffunc H2( Terminal of SCM-AE ) -> Element of INT = (id INT ) . (s . a1);
func c2 @ c1 -> Integer means :Def9: :: SCM_COMP:def 9
ex f being Function of TS SCM-AE , INT st
( it = f . nt1 & ( for t being Terminal of SCM-AE holds f . (root-tree t) = nt0 . t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of INT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = nt -Meaning_on xl,xr ) );
existence
ex b1 being Integer ex f being Function of TS SCM-AE , INT st
( b1 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = s . t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of INT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = nt -Meaning_on xl,xr ) )
proof end;
uniqueness
for b1, b2 being Integer st ex f being Function of TS SCM-AE , INT st
( b1 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = s . t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of INT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = nt -Meaning_on xl,xr ) ) & ex f being Function of TS SCM-AE , INT st
( b2 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = s . t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of INT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = nt -Meaning_on xl,xr ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def9 defines @ SCM_COMP:def 9 :
for s being State of SCM
for term being bin-term
for b3 being Integer holds
( b3 = term @ s iff ex f being Function of TS SCM-AE , INT st
( b3 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = s . t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of INT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = nt -Meaning_on xl,xr ) ) );

theorem Th6: :: SCM_COMP:6
for s being State of SCM
for t being Terminal of SCM-AE holds (root-tree t) @ s = s . t
proof end;

theorem Th7: :: SCM_COMP:7
for s being State of SCM
for nt being NonTerminal of SCM-AE
for tl, tr being bin-term holds (nt -tree tl,tr) @ s = nt -Meaning_on (tl @ s),(tr @ s)
proof end;

theorem Th8: :: SCM_COMP:8
for s being State of SCM
for tl, tr being bin-term holds
( (tl + tr) @ s = (tl @ s) + (tr @ s) & (tl - tr) @ s = (tl @ s) - (tr @ s) & (tl * tr) @ s = (tl @ s) * (tr @ s) & (tl div tr) @ s = (tl @ s) div (tr @ s) & (tl mod tr) @ s = (tl @ s) mod (tr @ s) )
proof end;

definition
let nt be NonTerminal of SCM-AE ;
let n be Element of NAT ;
func Selfwork c1,c2 -> Element of the Instructions of SCM * equals :Def10: :: SCM_COMP:def 10
<*(AddTo (dl. nt1),(dl. (nt1 + 1)))*> if nt0 = [0,0]
<*(SubFrom (dl. nt1),(dl. (nt1 + 1)))*> if nt0 = [0,1]
<*(MultBy (dl. nt1),(dl. (nt1 + 1)))*> if nt0 = [0,2]
<*(Divide (dl. nt1),(dl. (nt1 + 1)))*> if nt0 = [0,3]
<*(Divide (dl. nt1),(dl. (nt1 + 1))),((dl. nt1) := (dl. (nt1 + 1)))*> if nt0 = [0,4]
;
coherence
( ( nt = [0,0] implies <*(AddTo (dl. n),(dl. (n + 1)))*> is Element of the Instructions of SCM * ) & ( nt = [0,1] implies <*(SubFrom (dl. n),(dl. (n + 1)))*> is Element of the Instructions of SCM * ) & ( nt = [0,2] implies <*(MultBy (dl. n),(dl. (n + 1)))*> is Element of the Instructions of SCM * ) & ( nt = [0,3] implies <*(Divide (dl. n),(dl. (n + 1)))*> is Element of the Instructions of SCM * ) & ( nt = [0,4] implies <*(Divide (dl. n),(dl. (n + 1))),((dl. n) := (dl. (n + 1)))*> is Element of the Instructions of SCM * ) )
;
consistency
for b1 being Element of the Instructions of SCM * holds
( ( nt = [0,0] & nt = [0,1] implies ( b1 = <*(AddTo (dl. n),(dl. (n + 1)))*> iff b1 = <*(SubFrom (dl. n),(dl. (n + 1)))*> ) ) & ( nt = [0,0] & nt = [0,2] implies ( b1 = <*(AddTo (dl. n),(dl. (n + 1)))*> iff b1 = <*(MultBy (dl. n),(dl. (n + 1)))*> ) ) & ( nt = [0,0] & nt = [0,3] implies ( b1 = <*(AddTo (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1)))*> ) ) & ( nt = [0,0] & nt = [0,4] implies ( b1 = <*(AddTo (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1))),((dl. n) := (dl. (n + 1)))*> ) ) & ( nt = [0,1] & nt = [0,2] implies ( b1 = <*(SubFrom (dl. n),(dl. (n + 1)))*> iff b1 = <*(MultBy (dl. n),(dl. (n + 1)))*> ) ) & ( nt = [0,1] & nt = [0,3] implies ( b1 = <*(SubFrom (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1)))*> ) ) & ( nt = [0,1] & nt = [0,4] implies ( b1 = <*(SubFrom (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1))),((dl. n) := (dl. (n + 1)))*> ) ) & ( nt = [0,2] & nt = [0,3] implies ( b1 = <*(MultBy (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1)))*> ) ) & ( nt = [0,2] & nt = [0,4] implies ( b1 = <*(MultBy (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1))),((dl. n) := (dl. (n + 1)))*> ) ) & ( nt = [0,3] & nt = [0,4] implies ( b1 = <*(Divide (dl. n),(dl. (n + 1)))*> iff b1 = <*(Divide (dl. n),(dl. (n + 1))),((dl. n) := (dl. (n + 1)))*> ) ) )
proof end;
end;

:: deftheorem Def10 defines Selfwork SCM_COMP:def 10 :
for nt being NonTerminal of SCM-AE
for n being Element of NAT holds
( ( nt = [0,0] implies Selfwork nt,n = <*(AddTo (dl. n),(dl. (n + 1)))*> ) & ( nt = [0,1] implies Selfwork nt,n = <*(SubFrom (dl. n),(dl. (n + 1)))*> ) & ( nt = [0,2] implies Selfwork nt,n = <*(MultBy (dl. n),(dl. (n + 1)))*> ) & ( nt = [0,3] implies Selfwork nt,n = <*(Divide (dl. n),(dl. (n + 1)))*> ) & ( nt = [0,4] implies Selfwork nt,n = <*(Divide (dl. n),(dl. (n + 1))),((dl. n) := (dl. (n + 1)))*> ) );

definition
let term be bin-term;
let aux be Element of NAT ;
deffunc H1( Terminal of SCM-AE , Element of NAT ) -> Element of the Instructions of SCM * = <*((dl. a2) := (@ a1))*>;
deffunc H2( NonTerminal of SCM-AE , Function of NAT ,the Instructions of SCM * , Function of NAT ,the Instructions of SCM * , Element of NAT ) -> Element of the Instructions of SCM * = ((a2 . a4) ^ (a3 . (a4 + 1))) ^ (Selfwork a1,a4);
func SCM-Compile c1,c2 -> FinSequence of the Instructions of SCM means :Def11: :: SCM_COMP:def 11
ex f being Function of TS SCM-AE , Funcs NAT ,(the Instructions of SCM * ) st
( it = (f . nt0) . nt1 & ( for t being Terminal of SCM-AE ex g being Function of NAT ,the Instructions of SCM * st
( g = f . (root-tree t) & ( for n being Element of NAT holds g . n = <*((dl. n) := (@ t))*> ) ) ) & ( for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
ex g, f1, f2 being Function of NAT ,the Instructions of SCM * st
( g = f . (nt -tree t1,t2) & f1 = f . t1 & f2 = f . t2 & ( for n being Element of NAT holds g . n = ((f1 . n) ^ (f2 . (n + 1))) ^ (Selfwork nt,n) ) ) ) );
existence
ex b1 being FinSequence of the Instructions of SCM ex f being Function of TS SCM-AE , Funcs NAT ,(the Instructions of SCM * ) st
( b1 = (f . term) . aux & ( for t being Terminal of SCM-AE ex g being Function of NAT ,the Instructions of SCM * st
( g = f . (root-tree t) & ( for n being Element of NAT holds g . n = <*((dl. n) := (@ t))*> ) ) ) & ( for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
ex g, f1, f2 being Function of NAT ,the Instructions of SCM * st
( g = f . (nt -tree t1,t2) & f1 = f . t1 & f2 = f . t2 & ( for n being Element of NAT holds g . n = ((f1 . n) ^ (f2 . (n + 1))) ^ (Selfwork nt,n) ) ) ) )
proof end;
uniqueness
for b1, b2 being FinSequence of the Instructions of SCM st ex f being Function of TS SCM-AE , Funcs NAT ,(the Instructions of SCM * ) st
( b1 = (f . term) . aux & ( for t being Terminal of SCM-AE ex g being Function of NAT ,the Instructions of SCM * st
( g = f . (root-tree t) & ( for n being Element of NAT holds g . n = <*((dl. n) := (@ t))*> ) ) ) & ( for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
ex g, f1, f2 being Function of NAT ,the Instructions of SCM * st
( g = f . (nt -tree t1,t2) & f1 = f . t1 & f2 = f . t2 & ( for n being Element of NAT holds g . n = ((f1 . n) ^ (f2 . (n + 1))) ^ (Selfwork nt,n) ) ) ) ) & ex f being Function of TS SCM-AE , Funcs NAT ,(the Instructions of SCM * ) st
( b2 = (f . term) . aux & ( for t being Terminal of SCM-AE ex g being Function of NAT ,the Instructions of SCM * st
( g = f . (root-tree t) & ( for n being Element of NAT holds g . n = <*((dl. n) := (@ t))*> ) ) ) & ( for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
ex g, f1, f2 being Function of NAT ,the Instructions of SCM * st
( g = f . (nt -tree t1,t2) & f1 = f . t1 & f2 = f . t2 & ( for n being Element of NAT holds g . n = ((f1 . n) ^ (f2 . (n + 1))) ^ (Selfwork nt,n) ) ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def11 defines SCM-Compile SCM_COMP:def 11 :
for term being bin-term
for aux being Element of NAT
for b3 being FinSequence of the Instructions of SCM holds
( b3 = SCM-Compile term,aux iff ex f being Function of TS SCM-AE , Funcs NAT ,(the Instructions of SCM * ) st
( b3 = (f . term) . aux & ( for t being Terminal of SCM-AE ex g being Function of NAT ,the Instructions of SCM * st
( g = f . (root-tree t) & ( for n being Element of NAT holds g . n = <*((dl. n) := (@ t))*> ) ) ) & ( for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
ex g, f1, f2 being Function of NAT ,the Instructions of SCM * st
( g = f . (nt -tree t1,t2) & f1 = f . t1 & f2 = f . t2 & ( for n being Element of NAT holds g . n = ((f1 . n) ^ (f2 . (n + 1))) ^ (Selfwork nt,n) ) ) ) ) );

consider term' being bin-term, aux' being Element of NAT ;

consider f being Function of TS SCM-AE , Funcs NAT ,(the Instructions of SCM * ) such that
SCM-Compile term',aux' = (f . term') . aux' and
Lemma139: ( ( for t being Terminal of SCM-AE ex g being Function of NAT ,the Instructions of SCM * st
( g = f . (root-tree t) & ( for n being Element of NAT holds g . n = <*((dl. n) := (@ t))*> ) ) ) & ( for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
ex g, f1, f2 being Function of NAT ,the Instructions of SCM * st
( g = f . (nt -tree t1,t2) & f1 = f . t1 & f2 = f . t2 & ( for n being Element of NAT holds g . n = ((f1 . n) ^ (f2 . (n + 1))) ^ (Selfwork nt,n) ) ) ) ) by ;

theorem Th9: :: SCM_COMP:9
for t being Terminal of SCM-AE
for n being Element of NAT holds SCM-Compile (root-tree t),n = <*((dl. n) := (@ t))*>
proof end;

theorem Th10: :: SCM_COMP:10
for nt being NonTerminal of SCM-AE
for t1, t2 being bin-term
for n being Element of NAT
for rtl, rtr being Symbol of SCM-AE st rtl = root-label t1 & rtr = root-label t2 & nt ==> <*rtl,rtr*> holds
SCM-Compile (nt -tree t1,t2),n = ((SCM-Compile t1,n) ^ (SCM-Compile t2,(n + 1))) ^ (Selfwork nt,n)
proof end;

definition
let t be Terminal of SCM-AE ;
func d". c1 -> Element of NAT means :Def12: :: SCM_COMP:def 12
dl. it = nt0;
existence
ex b1 being Element of NAT st dl. b1 = t
proof end;
uniqueness
for b1, b2 being Element of NAT st dl. b1 = t & dl. b2 = t holds
b1 = b2
by AMI_3:52;
end;

:: deftheorem Def12 defines d". SCM_COMP:def 12 :
for t being Terminal of SCM-AE
for b2 being Element of NAT holds
( b2 = d". t iff dl. b2 = t );

definition
let term be bin-term;
deffunc H1( NonTerminal of SCM-AE , set , set , Element of NAT , Element of NAT ) -> Element of NAT = max a4,a5;
deffunc H2( Terminal of SCM-AE ) -> Element of NAT = d". a1;
func max_Data-Loc_in c1 -> Element of NAT means :Def13: :: SCM_COMP:def 13
ex f being Function of TS SCM-AE , NAT st
( it = f . nt0 & ( for t being Terminal of SCM-AE holds f . (root-tree t) = d". t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of NAT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = max xl,xr ) );
existence
ex b1 being Element of NAT ex f being Function of TS SCM-AE , NAT st
( b1 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = d". t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of NAT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = max xl,xr ) )
proof end;
uniqueness
for b1, b2 being Element of NAT st ex f being Function of TS SCM-AE , NAT st
( b1 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = d". t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of NAT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = max xl,xr ) ) & ex f being Function of TS SCM-AE , NAT st
( b2 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = d". t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of NAT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = max xl,xr ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def13 defines max_Data-Loc_in SCM_COMP:def 13 :
for term being bin-term
for b2 being Element of NAT holds
( b2 = max_Data-Loc_in term iff ex f being Function of TS SCM-AE , NAT st
( b2 = f . term & ( for t being Terminal of SCM-AE holds f . (root-tree t) = d". t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of NAT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = max xl,xr ) ) );

consider Term being bin-term;

consider f being Function of TS SCM-AE , NAT such that
max_Data-Loc_in Term = f . Term and
Lemma148: ( ( for t being Terminal of SCM-AE holds f . (root-tree t) = d". t ) & ( for nt being NonTerminal of SCM-AE
for tl, tr being bin-term
for rtl, rtr being Symbol of SCM-AE st rtl = root-label tl & rtr = root-label tr & nt ==> <*rtl,rtr*> holds
for xl, xr being Element of NAT st xl = f . tl & xr = f . tr holds
f . (nt -tree tl,tr) = max xl,xr ) ) by ;

theorem Th11: :: SCM_COMP:11
for t being Terminal of SCM-AE holds max_Data-Loc_in (root-tree t) = d". t
proof end;

Lemma150: NonTerminals SCM-AE = [:1,5:]
by ;

theorem Th12: :: SCM_COMP:12
for nt being NonTerminal of SCM-AE
for tl, tr being bin-term holds max_Data-Loc_in (nt -tree tl,tr) = max (max_Data-Loc_in tl),(max_Data-Loc_in tr)
proof end;

defpred S1[ bin-term] means for s1, s2 being State of SCM st ( for dn being Element of NAT st dn <= max_Data-Loc_in a1 holds
s1 . (dl. dn) = s2 . (dl. dn) ) holds
a1 @ s1 = a1 @ s2;

E155: now
let s be Terminal of SCM-AE ;
thus S1[ root-tree s]
proof
let s1 be State of SCM , s2 be State of SCM ;
assume E34: for dn being Element of NAT st dn <= max_Data-Loc_in (root-tree s) holds
s1 . (dl. dn) = s2 . (dl. dn) ;
d". s <= max_Data-Loc_in (root-tree s) by ;
then E35: ( s1 . (dl. (d". s)) = s2 . (dl. (d". s)) & s1 . s = s1 . (dl. (d". s)) & s2 . s = s2 . (dl. (d". s)) ) by , ;
(root-tree s) @ s1 = s1 . s by ;
hence (root-tree s) @ s1 = (root-tree s) @ s2 by , ;
end;
end;

E156: now
let nt be NonTerminal of SCM-AE ;
let tl be Element of TS SCM-AE , tr be Element of TS SCM-AE ;
assume E34: ( nt ==> <*(root-label tl),(root-label tr)*> & S1[tl] & S1[tr] ) ;
thus S1[nt -tree tl,tr]
proof
let s1 be State of SCM , s2 be State of SCM ;
assume E35: for dn being Element of NAT st dn <= max_Data-Loc_in (nt -tree tl,tr) holds
s1 . (dl. dn) = s2 . (dl. dn) ;
now
let dn be Element of NAT ;
assume E36: dn <= max_Data-Loc_in tl ;
set ml = max_Data-Loc_in tl;
set mr = max_Data-Loc_in tr;
max_Data-Loc_in tl <= max (max_Data-Loc_in tl),(max_Data-Loc_in tr) by SQUARE_1:46;
then dn <= max (max_Data-Loc_in tl),(max_Data-Loc_in tr) by , XREAL_1:2;
then dn <= max_Data-Loc_in (nt -tree tl,tr) by ;
hence s1 . (dl. dn) = s2 . (dl. dn) by Th1;
end;
then E38: tl @ s1 = tl @ s2 by ;
now
let dn be Element of NAT ;
assume E39: dn <= max_Data-Loc_in tr ;
set ml = max_Data-Loc_in tl;
set mr = max_Data-Loc_in tr;
max_Data-Loc_in tr <= max (max_Data-Loc_in tl),(max_Data-Loc_in tr) by SQUARE_1:46;
then dn <= max (max_Data-Loc_in tl),(max_Data-Loc_in tr) by , XREAL_1:2;
then dn <= max_Data-Loc_in (nt -tree tl,tr) by ;
hence s1 . (dl. dn) = s2 . (dl. dn) by Th1;
end;
then E40: tr @ s1 = tr @ s2 by ;
(nt -tree tl,tr) @ s1 = nt -Meaning_on (tl @ s1),(tr @ s1) by ;
hence (nt -tree tl,tr) @ s1 = (nt -tree tl,tr) @ s2 by , , ;
end;
end;

theorem Th13: :: SCM_COMP:13
for term being bin-term
for s1, s2 being State of SCM st ( for dn being Element of NAT st dn <= max_Data-Loc_in term holds
s1 . (dl. dn) = s2 . (dl. dn) ) holds
term @ s1 = term @ s2
proof end;

set D = <*> INT ;

defpred S2[ bin-term] means for aux, n, k being Element of NAT
for s being State-consisting of n,n,k, SCM-Compile a1,aux, <*> INT st aux > max_Data-Loc_in a1 holds
ex i being Element of NAT ex u being State of SCM st
( u = (Computation s) . (i + 1) & i + 1 = len (SCM-Compile a1,aux) & IC ((Computation s) . i) = il. (n + i) & IC u = il. (n + (i + 1)) & u . (dl. aux) = a1 @ s & ( for dn being Element of NAT st dn < aux holds
s . (dl. dn) = u . (dl. dn) ) );

theorem Th14: :: SCM_COMP:14
for term being bin-term
for aux, n, k being Element of NAT
for s being State-consisting of n,n,k, SCM-Compile term,aux, <*> INT st aux > max_Data-Loc_in term holds
ex i being Element of NAT ex u being State of SCM st
( u = (Computation s) . (i + 1) & i + 1 = len (SCM-Compile term,aux) & IC ((Computation s) . i) = il. (n + i) & IC u = il. (n + (i + 1)) & u . (dl. aux) = term @ s & ( for dn being Element of NAT st dn < aux holds
s . (dl. dn) = u . (dl. dn) ) )
proof end;

theorem Th15: :: SCM_COMP:15
for term being bin-term
for aux, n, k being Element of NAT
for s being State-consisting of n,n,k,(SCM-Compile term,aux) ^ <*(halt SCM )*>, <*> INT st aux > max_Data-Loc_in term holds
( s is halting & (Result s) . (dl. aux) = term @ s & Complexity s = len (SCM-Compile term,aux) )
proof end;