begin
:: deftheorem Def1 defines ==> LANG1:def 1 :
for G being non empty DTConstrStr
for s being Symbol of G
for n being FinSequence holds
( s ==> n iff [s,n] in the Rules of G );
:: deftheorem defines Terminals LANG1:def 2 :
for G being non empty DTConstrStr holds Terminals G = { s where s is Symbol of G : for n being FinSequence holds not s ==> n } ;
:: deftheorem defines NonTerminals LANG1:def 3 :
for G being non empty DTConstrStr holds NonTerminals G = { s where s is Symbol of G : ex n being FinSequence st s ==> n } ;
theorem
:: deftheorem defines ==> LANG1:def 4 :
for G being non empty DTConstrStr
for n, m being String of G holds
( n ==> m iff ex n1, n2, n3 being String of G ex s being Symbol of G st
( n = (n1 ^ <*s*>) ^ n2 & m = (n1 ^ n3) ^ n2 & s ==> n3 ) );
theorem
theorem Th3:
theorem Th4:
theorem Th5:
:: deftheorem Def5 defines is_derivable_from LANG1:def 5 :
for G being non empty DTConstrStr
for n, m being String of G holds
( m is_derivable_from n iff ex p being FinSequence st
( len p >= 1 & p . 1 = n & p . (len p) = m & ( for i being Element of NAT st i >= 1 & i < len p holds
ex a, b being String of G st
( p . i = a & p . (i + 1) = b & a ==> b ) ) ) );
theorem Th6:
theorem Th7:
theorem Th8:
:: deftheorem defines Lang LANG1:def 6 :
for G being non empty GrammarStr holds Lang G = { a where a is Element of the carrier of G * : ( rng a c= Terminals G & a is_derivable_from <* the InitialSym of G*> ) } ;
theorem
definition
let a be
set ;
func EmptyGrammar a -> strict GrammarStr means :
Def7:
( the
carrier of
it = {a} & the
Rules of
it = {[a,{}]} );
existence
ex b1 being strict GrammarStr st
( the carrier of b1 = {a} & the Rules of b1 = {[a,{}]} )
uniqueness
for b1, b2 being strict GrammarStr st the carrier of b1 = {a} & the Rules of b1 = {[a,{}]} & the carrier of b2 = {a} & the Rules of b2 = {[a,{}]} holds
b1 = b2
let b be
set ;
func SingleGrammar (
a,
b)
-> strict GrammarStr means :
Def8:
( the
carrier of
it = {a,b} & the
InitialSym of
it = a & the
Rules of
it = {[a,<*b*>]} );
existence
ex b1 being strict GrammarStr st
( the carrier of b1 = {a,b} & the InitialSym of b1 = a & the Rules of b1 = {[a,<*b*>]} )
uniqueness
for b1, b2 being strict GrammarStr st the carrier of b1 = {a,b} & the InitialSym of b1 = a & the Rules of b1 = {[a,<*b*>]} & the carrier of b2 = {a,b} & the InitialSym of b2 = a & the Rules of b2 = {[a,<*b*>]} holds
b1 = b2
;
func IterGrammar (
a,
b)
-> strict GrammarStr means :
Def9:
( the
carrier of
it = {a,b} & the
InitialSym of
it = a & the
Rules of
it = {[a,<*b,a*>],[a,{}]} );
existence
ex b1 being strict GrammarStr st
( the carrier of b1 = {a,b} & the InitialSym of b1 = a & the Rules of b1 = {[a,<*b,a*>],[a,{}]} )
uniqueness
for b1, b2 being strict GrammarStr st the carrier of b1 = {a,b} & the InitialSym of b1 = a & the Rules of b1 = {[a,<*b,a*>],[a,{}]} & the carrier of b2 = {a,b} & the InitialSym of b2 = a & the Rules of b2 = {[a,<*b,a*>],[a,{}]} holds
b1 = b2
;
end;
:: deftheorem Def7 defines EmptyGrammar LANG1:def 7 :
for a being set
for b2 being strict GrammarStr holds
( b2 = EmptyGrammar a iff ( the carrier of b2 = {a} & the Rules of b2 = {[a,{}]} ) );
:: deftheorem Def8 defines SingleGrammar LANG1:def 8 :
for a, b being set
for b3 being strict GrammarStr holds
( b3 = SingleGrammar (a,b) iff ( the carrier of b3 = {a,b} & the InitialSym of b3 = a & the Rules of b3 = {[a,<*b*>]} ) );
:: deftheorem Def9 defines IterGrammar LANG1:def 9 :
for a, b being set
for b3 being strict GrammarStr holds
( b3 = IterGrammar (a,b) iff ( the carrier of b3 = {a,b} & the InitialSym of b3 = a & the Rules of b3 = {[a,<*b,a*>],[a,{}]} ) );
definition
let D be non
empty set ;
func TotalGrammar D -> strict GrammarStr means :
Def10:
( the
carrier of
it = succ D & the
InitialSym of
it = D & the
Rules of
it = { [D,<*d,D*>] where d is Element of D : d = d } \/ {[D,{}]} );
existence
ex b1 being strict GrammarStr st
( the carrier of b1 = succ D & the InitialSym of b1 = D & the Rules of b1 = { [D,<*d,D*>] where d is Element of D : d = d } \/ {[D,{}]} )
uniqueness
for b1, b2 being strict GrammarStr st the carrier of b1 = succ D & the InitialSym of b1 = D & the Rules of b1 = { [D,<*d,D*>] where d is Element of D : d = d } \/ {[D,{}]} & the carrier of b2 = succ D & the InitialSym of b2 = D & the Rules of b2 = { [D,<*d,D*>] where d is Element of D : d = d } \/ {[D,{}]} holds
b1 = b2
;
end;
:: deftheorem Def10 defines TotalGrammar LANG1:def 10 :
for D being non empty set
for b2 being strict GrammarStr holds
( b2 = TotalGrammar D iff ( the carrier of b2 = succ D & the InitialSym of b2 = D & the Rules of b2 = { [D,<*d,D*>] where d is Element of D : d = d } \/ {[D,{}]} ) );
theorem Th10:
theorem Th11:
theorem Th12:
theorem
theorem Th14:
theorem
theorem Th16:
theorem
:: deftheorem Def11 defines efective LANG1:def 11 :
for IT being non empty GrammarStr holds
( IT is efective iff ( not Lang IT is empty & the InitialSym of IT in NonTerminals IT & ( for s being Symbol of IT st s in Terminals IT holds
ex p being String of IT st
( p in Lang IT & s in rng p ) ) ) );
:: deftheorem defines finite LANG1:def 12 :
for IT being GrammarStr holds
( IT is finite iff the Rules of IT is finite );
:: deftheorem LANG1:def 13 :
canceled;
:: deftheorem defines * LANG1:def 14 :
for X, Y being non empty set
for f being Function of X,Y
for b4 being Function of (X *),(Y *) holds
( b4 = f * iff for p being Element of X * holds b4 . p = f * p );
theorem
:: deftheorem defines . LANG1:def 15 :
for G being non empty GrammarStr
for X being non empty set
for f being Function of the carrier of G,X holds G . f = GrammarStr(# X,(f . the InitialSym of G),(((f ~) * the Rules of G) * (f *)) #);