:: FREEALG semantic presentation

definition
let IT be set ;
attr a1 is missing_with_Nat means :Def1: :: FREEALG:def 1
IT misses NAT ;
end;

:: deftheorem Def1 defines missing_with_Nat FREEALG:def 1 :
for IT being set holds
( IT is missing_with_Nat iff IT misses NAT );

registration
cluster non empty missing_with_Nat set ;
existence
ex b1 being set st
( not b1 is empty & b1 is missing_with_Nat )
proof end;
end;

Lemma33: ( not 0 in rng <*1*> & 0 in rng <*0*> )
proof end;

notation
let IT be Relation;
antonym with_zero c1 for non-empty c1;
synonym without_zero c1 for non-empty c1;
end;

definition
let IT be Relation;
redefine attr a1 is non-empty means :Def2: :: FREEALG:def 2
not 0 in rng IT;
compatibility
( not IT is with_zero iff not 0 in rng IT )
by RELAT_1:def 9;
end;

:: deftheorem Def2 defines with_zero FREEALG:def 2 :
for IT being Relation holds
( not IT is with_zero iff not 0 in rng IT );

registration
cluster non empty with_zero FinSequence of NAT ;
existence
ex b1 being FinSequence of NAT st
( not b1 is empty & b1 is with_zero )
proof end;
cluster non empty without_zero FinSequence of NAT ;
existence
ex b1 being FinSequence of NAT st
( not b1 is empty & not b1 is with_zero )
proof end;
end;

definition
let U1 be Universal_Algebra;
let n be Element of NAT ;
assume E32: n in dom the charact of U1 ;
canceled;
func oper c2,c1 -> operation of a1 equals :Def4: :: FREEALG:def 4
the charact of U1 . n;
coherence
the charact of U1 . n is operation of U1
by , FUNCT_1:def 5;
end;

:: deftheorem Def3 FREEALG:def 3 :
canceled;

:: deftheorem Def4 defines oper FREEALG:def 4 :
for U1 being Universal_Algebra
for n being Element of NAT st n in dom the charact of U1 holds
oper n,U1 = the charact of U1 . n;

definition
let U0 be Universal_Algebra;
mode GeneratorSet of c1 -> Subset of a1 means :Def5: :: FREEALG:def 5
the carrier of (GenUnivAlg it) = the carrier of U0;
existence
ex b1 being Subset of U0 st the carrier of (GenUnivAlg b1) = the carrier of U0
proof end;
end;

:: deftheorem Def5 defines GeneratorSet FREEALG:def 5 :
for U0 being Universal_Algebra
for b2 being Subset of U0 holds
( b2 is GeneratorSet of U0 iff the carrier of (GenUnivAlg b2) = the carrier of U0 );

definition
let U0 be Universal_Algebra;
let IT be GeneratorSet of U0;
attr a2 is free means :Def6: :: FREEALG:def 6
for U1 being Universal_Algebra st U0,U1 are_similar holds
for f being Function of IT,the carrier of U1 ex h being Function of U0,U1 st
( h is_homomorphism U0,U1 & h | IT = f );
end;

:: deftheorem Def6 defines free FREEALG:def 6 :
for U0 being Universal_Algebra
for IT being GeneratorSet of U0 holds
( IT is free iff for U1 being Universal_Algebra st U0,U1 are_similar holds
for f being Function of IT,the carrier of U1 ex h being Function of U0,U1 st
( h is_homomorphism U0,U1 & h | IT = f ) );

definition
let IT be Universal_Algebra;
attr a1 is free means :Def7: :: FREEALG:def 7
ex G being GeneratorSet of IT st G is free;
end;

:: deftheorem Def7 defines free FREEALG:def 7 :
for IT being Universal_Algebra holds
( IT is free iff ex G being GeneratorSet of IT st G is free );

registration
cluster strict free UAStr ;
existence
ex b1 being Universal_Algebra st
( b1 is free & b1 is strict )
proof end;
end;

registration
let U0 be free Universal_Algebra;
cluster free GeneratorSet of a1;
existence
ex b1 being GeneratorSet of U0 st b1 is free
by ;
end;

theorem Th1: :: FREEALG:1
for U0 being strict Universal_Algebra
for A being Subset of U0 holds
( A is GeneratorSet of U0 iff GenUnivAlg A = U0 )
proof end;

definition
let f be non empty FinSequence of NAT ;
let X be set ;
func REL c1,c2 -> Relation of (dom a1) \/ a2,((dom a1) \/ a2) * means :Def8: :: FREEALG:def 8
for a being Element of (dom f) \/ X
for b being Element of ((dom f) \/ X) * holds
( [a,b] in it iff ( a in dom f & f . a = len b ) );
existence
ex b1 being Relation of (dom f) \/ X,((dom f) \/ X) * st
for a being Element of (dom f) \/ X
for b being Element of ((dom f) \/ X) * holds
( [a,b] in b1 iff ( a in dom f & f . a = len b ) )
proof end;
uniqueness
for b1, b2 being Relation of (dom f) \/ X,((dom f) \/ X) * st ( for a being Element of (dom f) \/ X
for b being Element of ((dom f) \/ X) * holds
( [a,b] in b1 iff ( a in dom f & f . a = len b ) ) ) & ( for a being Element of (dom f) \/ X
for b being Element of ((dom f) \/ X) * holds
( [a,b] in b2 iff ( a in dom f & f . a = len b ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def8 defines REL FREEALG:def 8 :
for f being non empty FinSequence of NAT
for X being set
for b3 being Relation of (dom f) \/ X,((dom f) \/ X) * holds
( b3 = REL f,X iff for a being Element of (dom f) \/ X
for b being Element of ((dom f) \/ X) * holds
( [a,b] in b3 iff ( a in dom f & f . a = len b ) ) );

definition
let f be non empty FinSequence of NAT ;
let X be set ;
func DTConUA c1,c2 -> strict DTConstrStr equals :: FREEALG:def 9
DTConstrStr(# ((dom f) \/ X),(REL f,X) #);
correctness
coherence
DTConstrStr(# ((dom f) \/ X),(REL f,X) #) is strict DTConstrStr
;
;
end;

:: deftheorem Def9 defines DTConUA FREEALG:def 9 :
for f being non empty FinSequence of NAT
for X being set holds DTConUA f,X = DTConstrStr(# ((dom f) \/ X),(REL f,X) #);

registration
let f be non empty FinSequence of NAT ;
let X be set ;
cluster DTConUA a1,a2 -> non empty strict ;
coherence
not DTConUA f,X is empty
;
end;

theorem Th2: :: FREEALG:2
for f being non empty FinSequence of NAT
for X being set holds
( Terminals (DTConUA f,X) c= X & NonTerminals (DTConUA f,X) = dom f )
proof end;

theorem Th3: :: FREEALG:3
for f being non empty FinSequence of NAT
for X being missing_with_Nat set holds Terminals (DTConUA f,X) = X
proof end;

registration
let f be non empty FinSequence of NAT ;
let X be set ;
cluster DTConUA a1,a2 -> non empty strict with_nonterminals ;
coherence
DTConUA f,X is with_nonterminals
proof end;
end;

registration
let f be non empty with_zero FinSequence of NAT ;
let X be set ;
cluster DTConUA a1,a2 -> non empty strict with_nonterminals with_useful_nonterminals ;
coherence
( DTConUA f,X is with_nonterminals & DTConUA f,X is with_useful_nonterminals )
proof end;
end;

registration
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
cluster DTConUA a1,a2 -> non empty strict with_terminals with_nonterminals with_useful_nonterminals ;
coherence
( DTConUA f,D is with_terminals & DTConUA f,D is with_nonterminals & DTConUA f,D is with_useful_nonterminals )
proof end;
end;

definition
let f be non empty FinSequence of NAT ;
let X be set ;
let n be Nat;
assume E32: n in dom f ;
func Sym c3,c1,c2 -> Symbol of (DTConUA a1,a2) equals :Def10: :: FREEALG:def 10
n;
coherence
n is Symbol of (DTConUA f,X)
by , XBOOLE_0:def 2;
end;

:: deftheorem Def10 defines Sym FREEALG:def 10 :
for f being non empty FinSequence of NAT
for X being set
for n being Nat st n in dom f holds
Sym n,f,X = n;

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
let n be Nat;
assume E32: n in dom f ;
func FreeOpNSG c3,c1,c2 -> non empty homogeneous quasi_total PartFunc of (TS (DTConUA a1,a2)) * , TS (DTConUA a1,a2) means :Def11: :: FREEALG:def 11
( dom it = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom it holds
it . p = (Sym n,f,D) -tree p ) );
existence
ex b1 being non empty homogeneous quasi_total PartFunc of (TS (DTConUA f,D)) * , TS (DTConUA f,D) st
( dom b1 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b1 holds
b1 . p = (Sym n,f,D) -tree p ) )
proof end;
uniqueness
for b1, b2 being non empty homogeneous quasi_total PartFunc of (TS (DTConUA f,D)) * , TS (DTConUA f,D) st dom b1 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b1 holds
b1 . p = (Sym n,f,D) -tree p ) & dom b2 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b2 holds
b2 . p = (Sym n,f,D) -tree p ) holds
b1 = b2
proof end;
end;

:: deftheorem Def11 defines FreeOpNSG FREEALG:def 11 :
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set
for n being Nat st n in dom f holds
for b4 being non empty homogeneous quasi_total PartFunc of (TS (DTConUA f,D)) * , TS (DTConUA f,D) holds
( b4 = FreeOpNSG n,f,D iff ( dom b4 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b4 holds
b4 . p = (Sym n,f,D) -tree p ) ) );

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
func FreeOpSeqNSG c1,c2 -> PFuncFinSequence of (TS (DTConUA a1,a2)) means :Def12: :: FREEALG:def 12
( len it = len f & ( for n being Element of NAT st n in dom it holds
it . n = FreeOpNSG n,f,D ) );
existence
ex b1 being PFuncFinSequence of (TS (DTConUA f,D)) st
( len b1 = len f & ( for n being Element of NAT st n in dom b1 holds
b1 . n = FreeOpNSG n,f,D ) )
proof end;
uniqueness
for b1, b2 being PFuncFinSequence of (TS (DTConUA f,D)) st len b1 = len f & ( for n being Element of NAT st n in dom b1 holds
b1 . n = FreeOpNSG n,f,D ) & len b2 = len f & ( for n being Element of NAT st n in dom b2 holds
b2 . n = FreeOpNSG n,f,D ) holds
b1 = b2
proof end;
end;

:: deftheorem Def12 defines FreeOpSeqNSG FREEALG:def 12 :
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set
for b3 being PFuncFinSequence of (TS (DTConUA f,D)) holds
( b3 = FreeOpSeqNSG f,D iff ( len b3 = len f & ( for n being Element of NAT st n in dom b3 holds
b3 . n = FreeOpNSG n,f,D ) ) );

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
func FreeUnivAlgNSG c1,c2 -> strict Universal_Algebra equals :: FREEALG:def 13
UAStr(# (TS (DTConUA f,D)),(FreeOpSeqNSG f,D) #);
coherence
UAStr(# (TS (DTConUA f,D)),(FreeOpSeqNSG f,D) #) is strict Universal_Algebra
proof end;
end;

:: deftheorem Def13 defines FreeUnivAlgNSG FREEALG:def 13 :
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set holds FreeUnivAlgNSG f,D = UAStr(# (TS (DTConUA f,D)),(FreeOpSeqNSG f,D) #);

theorem Th4: :: FREEALG:4
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set holds signature (FreeUnivAlgNSG f,D) = f
proof end;

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
func FreeGenSetNSG c1,c2 -> Subset of (FreeUnivAlgNSG a1,a2) equals :: FREEALG:def 14
{ (root-tree s) where s is Symbol of (DTConUA f,D) : s in Terminals (DTConUA f,D) } ;
coherence
{ (root-tree s) where s is Symbol of (DTConUA f,D) : s in Terminals (DTConUA f,D) } is Subset of (FreeUnivAlgNSG f,D)
proof end;
end;

:: deftheorem Def14 defines FreeGenSetNSG FREEALG:def 14 :
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set holds FreeGenSetNSG f,D = { (root-tree s) where s is Symbol of (DTConUA f,D) : s in Terminals (DTConUA f,D) } ;

theorem Th5: :: FREEALG:5
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set holds not FreeGenSetNSG f,D is empty
proof end;

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
redefine func FreeGenSetNSG as FreeGenSetNSG c1,c2 -> GeneratorSet of FreeUnivAlgNSG a1,a2;
coherence
FreeGenSetNSG f,D is GeneratorSet of FreeUnivAlgNSG f,D
proof end;
end;

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
let C be non empty set ;
let s be Symbol of (DTConUA f,D);
let F be Function of FreeGenSetNSG f,D,C;
assume E32: s in Terminals (DTConUA f,D) ;
func pi c5,c4 -> Element of a3 equals :Def15: :: FREEALG:def 15
F . (root-tree s);
coherence
F . (root-tree s) is Element of C
proof end;
end;

:: deftheorem Def15 defines pi FREEALG:def 15 :
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set
for C being non empty set
for s being Symbol of (DTConUA f,D)
for F being Function of FreeGenSetNSG f,D,C st s in Terminals (DTConUA f,D) holds
pi F,s = F . (root-tree s);

definition
let f be non empty FinSequence of NAT ;
let D be set ;
let s be Symbol of (DTConUA f,D);
given p being FinSequence such that E32: s ==> p ;
func @ c3 -> Element of NAT equals :Def16: :: FREEALG:def 16
s;
coherence
s is Element of NAT
proof end;
end;

:: deftheorem Def16 defines @ FREEALG:def 16 :
for f being non empty FinSequence of NAT
for D being set
for s being Symbol of (DTConUA f,D) st ex p being FinSequence st s ==> p holds
@ s = s;

theorem Th6: :: FREEALG:6
for f being non empty FinSequence of NAT
for D being non empty missing_with_Nat set holds FreeGenSetNSG f,D is free
proof end;

registration
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
cluster FreeUnivAlgNSG a1,a2 -> strict free ;
coherence
FreeUnivAlgNSG f,D is free
proof end;
end;

definition
let f be non empty FinSequence of NAT ;
let D be non empty missing_with_Nat set ;
redefine func FreeGenSetNSG as FreeGenSetNSG c1,c2 -> free GeneratorSet of FreeUnivAlgNSG a1,a2;
coherence
FreeGenSetNSG f,D is free GeneratorSet of FreeUnivAlgNSG f,D
by ;
end;

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
let n be Nat;
assume E32: n in dom f ;
func FreeOpZAO c3,c1,c2 -> non empty homogeneous quasi_total PartFunc of (TS (DTConUA a1,a2)) * , TS (DTConUA a1,a2) means :Def17: :: FREEALG:def 17
( dom it = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom it holds
it . p = (Sym n,f,D) -tree p ) );
existence
ex b1 being non empty homogeneous quasi_total PartFunc of (TS (DTConUA f,D)) * , TS (DTConUA f,D) st
( dom b1 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b1 holds
b1 . p = (Sym n,f,D) -tree p ) )
proof end;
uniqueness
for b1, b2 being non empty homogeneous quasi_total PartFunc of (TS (DTConUA f,D)) * , TS (DTConUA f,D) st dom b1 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b1 holds
b1 . p = (Sym n,f,D) -tree p ) & dom b2 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b2 holds
b2 . p = (Sym n,f,D) -tree p ) holds
b1 = b2
proof end;
end;

:: deftheorem Def17 defines FreeOpZAO FREEALG:def 17 :
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set
for n being Nat st n in dom f holds
for b4 being non empty homogeneous quasi_total PartFunc of (TS (DTConUA f,D)) * , TS (DTConUA f,D) holds
( b4 = FreeOpZAO n,f,D iff ( dom b4 = (f /. n) -tuples_on (TS (DTConUA f,D)) & ( for p being FinSequence of TS (DTConUA f,D) st p in dom b4 holds
b4 . p = (Sym n,f,D) -tree p ) ) );

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
func FreeOpSeqZAO c1,c2 -> PFuncFinSequence of (TS (DTConUA a1,a2)) means :Def18: :: FREEALG:def 18
( len it = len f & ( for n being Element of NAT st n in dom it holds
it . n = FreeOpZAO n,f,D ) );
existence
ex b1 being PFuncFinSequence of (TS (DTConUA f,D)) st
( len b1 = len f & ( for n being Element of NAT st n in dom b1 holds
b1 . n = FreeOpZAO n,f,D ) )
proof end;
uniqueness
for b1, b2 being PFuncFinSequence of (TS (DTConUA f,D)) st len b1 = len f & ( for n being Element of NAT st n in dom b1 holds
b1 . n = FreeOpZAO n,f,D ) & len b2 = len f & ( for n being Element of NAT st n in dom b2 holds
b2 . n = FreeOpZAO n,f,D ) holds
b1 = b2
proof end;
end;

:: deftheorem Def18 defines FreeOpSeqZAO FREEALG:def 18 :
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set
for b3 being PFuncFinSequence of (TS (DTConUA f,D)) holds
( b3 = FreeOpSeqZAO f,D iff ( len b3 = len f & ( for n being Element of NAT st n in dom b3 holds
b3 . n = FreeOpZAO n,f,D ) ) );

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
func FreeUnivAlgZAO c1,c2 -> strict Universal_Algebra equals :: FREEALG:def 19
UAStr(# (TS (DTConUA f,D)),(FreeOpSeqZAO f,D) #);
coherence
UAStr(# (TS (DTConUA f,D)),(FreeOpSeqZAO f,D) #) is strict Universal_Algebra
proof end;
end;

:: deftheorem Def19 defines FreeUnivAlgZAO FREEALG:def 19 :
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set holds FreeUnivAlgZAO f,D = UAStr(# (TS (DTConUA f,D)),(FreeOpSeqZAO f,D) #);

theorem Th7: :: FREEALG:7
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set holds signature (FreeUnivAlgZAO f,D) = f
proof end;

theorem Th8: :: FREEALG:8
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set holds FreeUnivAlgZAO f,D is with_const_op
proof end;

theorem Th9: :: FREEALG:9
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set holds Constants (FreeUnivAlgZAO f,D) <> {}
proof end;

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
func FreeGenSetZAO c1,c2 -> Subset of (FreeUnivAlgZAO a1,a2) equals :: FREEALG:def 20
{ (root-tree s) where s is Symbol of (DTConUA f,D) : s in Terminals (DTConUA f,D) } ;
coherence
{ (root-tree s) where s is Symbol of (DTConUA f,D) : s in Terminals (DTConUA f,D) } is Subset of (FreeUnivAlgZAO f,D)
proof end;
end;

:: deftheorem Def20 defines FreeGenSetZAO FREEALG:def 20 :
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set holds FreeGenSetZAO f,D = { (root-tree s) where s is Symbol of (DTConUA f,D) : s in Terminals (DTConUA f,D) } ;

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
redefine func FreeGenSetZAO as FreeGenSetZAO c1,c2 -> GeneratorSet of FreeUnivAlgZAO a1,a2;
coherence
FreeGenSetZAO f,D is GeneratorSet of FreeUnivAlgZAO f,D
proof end;
end;

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
let C be non empty set ;
let s be Symbol of (DTConUA f,D);
let F be Function of FreeGenSetZAO f,D,C;
assume E32: s in Terminals (DTConUA f,D) ;
func pi c5,c4 -> Element of a3 equals :Def21: :: FREEALG:def 21
F . (root-tree s);
coherence
F . (root-tree s) is Element of C
proof end;
end;

:: deftheorem Def21 defines pi FREEALG:def 21 :
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set
for C being non empty set
for s being Symbol of (DTConUA f,D)
for F being Function of FreeGenSetZAO f,D,C st s in Terminals (DTConUA f,D) holds
pi F,s = F . (root-tree s);

theorem Th10: :: FREEALG:10
for f being non empty with_zero FinSequence of NAT
for D being missing_with_Nat set holds FreeGenSetZAO f,D is free
proof end;

registration
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
cluster FreeUnivAlgZAO a1,a2 -> strict free ;
coherence
FreeUnivAlgZAO f,D is free
proof end;
end;

definition
let f be non empty with_zero FinSequence of NAT ;
let D be missing_with_Nat set ;
redefine func FreeGenSetZAO as FreeGenSetZAO c1,c2 -> free GeneratorSet of FreeUnivAlgZAO a1,a2;
coherence
FreeGenSetZAO f,D is free GeneratorSet of FreeUnivAlgZAO f,D
by ;
end;

registration
cluster strict with_const_op free UAStr ;
existence
ex b1 being Universal_Algebra st
( b1 is strict & b1 is free & b1 is with_const_op )
proof end;
end;