begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
for
i,
j being
set for
b1,
b2 being
bag of
j for
b19,
b29 being
bag of
i st
b19 = b1 | i &
b29 = b2 | i &
b1 divides b2 holds
b19 divides b29
theorem Th5:
for
i,
j being
set for
b1,
b2 being
bag of
j for
b19,
b29 being
bag of
i st
b19 = b1 | i &
b29 = b2 | i holds
(
(b1 -' b2) | i = b19 -' b29 &
(b1 + b2) | i = b19 + b29 )
:: deftheorem Def1 defines bag_extend HILBASIS:def 1 :
for n, k being Nat
for b being bag of n
for b4 being Element of Bags (n + 1) holds
( b4 = b bag_extend k iff ( b4 | n = b & b4 . n = k ) );
theorem Th6:
theorem Th7:
:: deftheorem defines UnitBag HILBASIS:def 2 :
for X being set
for x being Element of X holds UnitBag x = (EmptyBag X) +* (x,1);
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
:: deftheorem defines 1_1 HILBASIS:def 3 :
for X being set
for x being Element of X
for L being non empty unital multLoopStr_0 holds 1_1 (x,L) = (0_ (X,L)) +* ((UnitBag x),(1_ L));
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
:: deftheorem defines minlen HILBASIS:def 4 :
for L being non empty right_complementable add-associative right_zeroed well-unital distributive doubleLoopStr
for I being non empty Subset of (Polynom-Ring L) holds minlen I = { x where x is Element of I : for x9, y9 being Polynomial of L st x9 = x & y9 in I holds
len x9 <= len y9 } ;
theorem Th17:
:: deftheorem Def5 defines monomial HILBASIS:def 5 :
for L being non empty right_complementable add-associative right_zeroed well-unital distributive doubleLoopStr
for n being Nat
for a being Element of L
for b4 being Polynomial of L holds
( b4 = monomial (a,n) iff for x being Nat holds
( ( x = n implies b4 . x = a ) & ( x <> n implies b4 . x = 0. L ) ) );
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
begin
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
definition
let R be non
empty non
trivial right_complementable Abelian add-associative right_zeroed associative commutative well-unital distributive doubleLoopStr ;
let n be
Element of
NAT ;
func upm (
n,
R)
-> Function of
(Polynom-Ring (Polynom-Ring (n,R))),
(Polynom-Ring ((n + 1),R)) means :
Def6:
for
p1 being
Polynomial of
(Polynom-Ring (n,R)) for
p2 being
Polynomial of
n,
R for
p3 being
Polynomial of
(n + 1),
R for
b being
bag of
n + 1 st
p3 = it . p1 &
p2 = p1 . (b . n) holds
p3 . b = p2 . (b | n);
existence
ex b1 being Function of (Polynom-Ring (Polynom-Ring (n,R))),(Polynom-Ring ((n + 1),R)) st
for p1 being Polynomial of (Polynom-Ring (n,R))
for p2 being Polynomial of n,R
for p3 being Polynomial of (n + 1),R
for b being bag of n + 1 st p3 = b1 . p1 & p2 = p1 . (b . n) holds
p3 . b = p2 . (b | n)
uniqueness
for b1, b2 being Function of (Polynom-Ring (Polynom-Ring (n,R))),(Polynom-Ring ((n + 1),R)) st ( for p1 being Polynomial of (Polynom-Ring (n,R))
for p2 being Polynomial of n,R
for p3 being Polynomial of (n + 1),R
for b being bag of n + 1 st p3 = b1 . p1 & p2 = p1 . (b . n) holds
p3 . b = p2 . (b | n) ) & ( for p1 being Polynomial of (Polynom-Ring (n,R))
for p2 being Polynomial of n,R
for p3 being Polynomial of (n + 1),R
for b being bag of n + 1 st p3 = b2 . p1 & p2 = p1 . (b . n) holds
p3 . b = p2 . (b | n) ) holds
b1 = b2
end;
:: deftheorem Def6 defines upm HILBASIS:def 6 :
for R being non empty non trivial right_complementable Abelian add-associative right_zeroed associative commutative well-unital distributive doubleLoopStr
for n being Element of NAT
for b3 being Function of (Polynom-Ring (Polynom-Ring (n,R))),(Polynom-Ring ((n + 1),R)) holds
( b3 = upm (n,R) iff for p1 being Polynomial of (Polynom-Ring (n,R))
for p2 being Polynomial of n,R
for p3 being Polynomial of (n + 1),R
for b being bag of n + 1 st p3 = b3 . p1 & p2 = p1 . (b . n) holds
p3 . b = p2 . (b | n) );
definition
let R be non
empty non
trivial right_complementable Abelian add-associative right_zeroed associative commutative well-unital distributive doubleLoopStr ;
let n be
Element of
NAT ;
func mpu (
n,
R)
-> Function of
(Polynom-Ring ((n + 1),R)),
(Polynom-Ring (Polynom-Ring (n,R))) means :
Def7:
for
p1 being
Polynomial of
(n + 1),
R for
p2 being
Polynomial of
n,
R for
p3 being
Polynomial of
(Polynom-Ring (n,R)) for
i being
Element of
NAT for
b being
bag of
n st
p3 = it . p1 &
p2 = p3 . i holds
p2 . b = p1 . (b bag_extend i);
existence
ex b1 being Function of (Polynom-Ring ((n + 1),R)),(Polynom-Ring (Polynom-Ring (n,R))) st
for p1 being Polynomial of (n + 1),R
for p2 being Polynomial of n,R
for p3 being Polynomial of (Polynom-Ring (n,R))
for i being Element of NAT
for b being bag of n st p3 = b1 . p1 & p2 = p3 . i holds
p2 . b = p1 . (b bag_extend i)
uniqueness
for b1, b2 being Function of (Polynom-Ring ((n + 1),R)),(Polynom-Ring (Polynom-Ring (n,R))) st ( for p1 being Polynomial of (n + 1),R
for p2 being Polynomial of n,R
for p3 being Polynomial of (Polynom-Ring (n,R))
for i being Element of NAT
for b being bag of n st p3 = b1 . p1 & p2 = p3 . i holds
p2 . b = p1 . (b bag_extend i) ) & ( for p1 being Polynomial of (n + 1),R
for p2 being Polynomial of n,R
for p3 being Polynomial of (Polynom-Ring (n,R))
for i being Element of NAT
for b being bag of n st p3 = b2 . p1 & p2 = p3 . i holds
p2 . b = p1 . (b bag_extend i) ) holds
b1 = b2
end;
:: deftheorem Def7 defines mpu HILBASIS:def 7 :
for R being non empty non trivial right_complementable Abelian add-associative right_zeroed associative commutative well-unital distributive doubleLoopStr
for n being Element of NAT
for b3 being Function of (Polynom-Ring ((n + 1),R)),(Polynom-Ring (Polynom-Ring (n,R))) holds
( b3 = mpu (n,R) iff for p1 being Polynomial of (n + 1),R
for p2 being Polynomial of n,R
for p3 being Polynomial of (Polynom-Ring (n,R))
for i being Element of NAT
for b being bag of n st p3 = b3 . p1 & p2 = p3 . i holds
p2 . b = p1 . (b bag_extend i) );
theorem Th30:
theorem Th31:
begin
theorem
canceled;
theorem Th33:
theorem Th34:
theorem
theorem