:: HILBASIS semantic presentation

theorem Th1: :: HILBASIS:1
for b1, b2 being FinSequence
for b3 being Function holds
not ( (rng b1) \/ (rng b2) c= dom b3 & ( for b4, b5 being FinSequence holds
not ( b4 = b3 * b1 & b5 = b3 * b2 & b3 * (b1 ^ b2) = b4 ^ b5 ) ) )
proof end;

theorem Th2: :: HILBASIS:2
for b1 being bag of 0 holds decomp b1 = <*<*{} ,{} *>*>
proof end;

theorem Th3: :: HILBASIS:3
for b1, b2 being Nat
for b3 being bag of b2 holds
( b1 <= b2 implies b3 | b1 is Element of Bags b1 )
proof end;

theorem Th4: :: HILBASIS:4
for b1, b2 being set
for b3, b4 being bag of b2
for b5, b6 being bag of b1 holds
( b5 = b3 | b1 & b6 = b4 | b1 & b3 divides b4 implies b5 divides b6 )
proof end;

theorem Th5: :: HILBASIS:5
for b1, b2 being set
for b3, b4 being bag of b2
for b5, b6 being bag of b1 holds
( b5 = b3 | b1 & b6 = b4 | b1 implies ( (b3 -' b4) | b1 = b5 -' b6 & (b3 + b4) | b1 = b5 + b6 ) )
proof end;

definition
let c1, c2 be Nat;
let c3 be bag of c1;
func c3 bag_extend c2 -> Element of Bags (a1 + 1) means :Def1: :: HILBASIS:def 1
( a4 | a1 = a3 & a4 . a1 = a2 );
existence
ex b1 being Element of Bags (c1 + 1) st
( b1 | c1 = c3 & b1 . c1 = c2 )
proof end;
uniqueness
for b1, b2 being Element of Bags (c1 + 1) holds
( b1 | c1 = c3 & b1 . c1 = c2 & b2 | c1 = c3 & b2 . c1 = c2 implies b1 = b2 )
proof end;
end;

:: deftheorem Def1 defines bag_extend HILBASIS:def 1 :
for b1, b2 being Nat
for b3 being bag of b1
for b4 being Element of Bags (b1 + 1) holds
( b4 = b3 bag_extend b2 iff ( b4 | b1 = b3 & b4 . b1 = b2 ) );

theorem Th6: :: HILBASIS:6
for b1 being Nat holds EmptyBag (b1 + 1) = (EmptyBag b1) bag_extend 0
proof end;

theorem Th7: :: HILBASIS:7
for b1 being Ordinal
for b2, b3 being bag of b1 holds
( b3 in rng (divisors b2) iff b3 divides b2 )
proof end;

definition
let c1 be set ;
let c2 be Element of c1;
func UnitBag c2 -> Element of Bags a1 equals :: HILBASIS:def 2
(EmptyBag a1) +* a2,1;
coherence
(EmptyBag c1) +* c2,1 is Element of Bags c1
by POLYNOM1:def 14;
end;

:: deftheorem Def2 defines UnitBag HILBASIS:def 2 :
for b1 being set
for b2 being Element of b1 holds UnitBag b2 = (EmptyBag b1) +* b2,1;

theorem Th8: :: HILBASIS:8
for b1 being non empty set
for b2 being Element of b1 holds support (UnitBag b2) = {b2}
proof end;

theorem Th9: :: HILBASIS:9
for b1 being non empty set
for b2 being Element of b1 holds
( (UnitBag b2) . b2 = 1 & ( for b3 being Element of b1 holds
( b2 <> b3 implies (UnitBag b2) . b3 = 0 ) ) )
proof end;

theorem Th10: :: HILBASIS:10
for b1 being non empty set
for b2, b3 being Element of b1 holds
( UnitBag b2 = UnitBag b3 implies b2 = b3 )
proof end;

theorem Th11: :: HILBASIS:11
for b1 being non empty Ordinal
for b2 being Element of b1
for b3 being non empty unital non trivial doubleLoopStr
for b4 being Function of b1,b3 holds eval (UnitBag b2),b4 = b4 . b2
proof end;

definition
let c1 be set ;
let c2 be Element of c1;
let c3 be non empty unital multLoopStr_0 ;
func 1_1 c2,c3 -> Series of a1,a3 equals :: HILBASIS:def 3
(0_ a1,a3) +* (UnitBag a2),(1. a3);
coherence
(0_ c1,c3) +* (UnitBag c2),(1. c3) is Series of c1,c3
;
end;

:: deftheorem Def3 defines 1_1 HILBASIS:def 3 :
for b1 being set
for b2 being Element of b1
for b3 being non empty unital multLoopStr_0 holds 1_1 b2,b3 = (0_ b1,b3) +* (UnitBag b2),(1. b3);

theorem Th12: :: HILBASIS:12
for b1 being set
for b2 being non empty unital non trivial doubleLoopStr
for b3 being Element of b1 holds
( (1_1 b3,b2) . (UnitBag b3) = 1. b2 & ( for b4 being bag of b1 holds
( b4 <> UnitBag b3 implies (1_1 b3,b2) . b4 = 0. b2 ) ) )
proof end;

theorem Th13: :: HILBASIS:13
for b1 being set
for b2 being Element of b1
for b3 being non empty add-associative right_zeroed right_complementable unital right-distributive non trivial doubleLoopStr holds Support (1_1 b2,b3) = {(UnitBag b2)}
proof end;

registration
let c1 be Ordinal;
let c2 be Element of c1;
let c3 be non empty add-associative right_zeroed right_complementable unital right-distributive non trivial doubleLoopStr ;
cluster 1_1 a2,a3 -> finite-Support ;
coherence
1_1 c2,c3 is finite-Support
proof end;
end;

theorem Th14: :: HILBASIS:14
for b1 being non empty add-associative right_zeroed right_complementable unital right-distributive non trivial doubleLoopStr
for b2 being non empty set
for b3, b4 being Element of b2 holds
( 1_1 b3,b1 = 1_1 b4,b1 implies b3 = b4 )
proof end;

theorem Th15: :: HILBASIS:15
for b1 being non empty add-associative right_zeroed right_complementable distributive doubleLoopStr
for b2 being Element of (Polynom-Ring b1)
for b3 being sequence of b1 holds
( b2 = b3 implies - b2 = - b3 )
proof end;

theorem Th16: :: HILBASIS:16
for b1 being non empty add-associative right_zeroed right_complementable distributive doubleLoopStr
for b2, b3 being Element of (Polynom-Ring b1)
for b4, b5 being sequence of b1 holds
( b2 = b4 & b3 = b5 implies b2 - b3 = b4 - b5 )
proof end;

definition
let c1 be non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr ;
let c2 be non empty Subset of (Polynom-Ring c1);
func minlen c2 -> non empty Subset of a2 equals :: HILBASIS:def 4
{ b1 where B is Element of a2 : for b1, b2 being Polynomial of a1 holds
( b2 = b1 & b3 in a2 implies len b2 <= len b3 )
}
;
coherence
{ b1 where B is Element of c2 : for b1, b2 being Polynomial of c1 holds
( b2 = b1 & b3 in c2 implies len b2 <= len b3 )
}
is non empty Subset of c2
proof end;
end;

:: deftheorem Def4 defines minlen HILBASIS:def 4 :
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2 being non empty Subset of (Polynom-Ring b1) holds minlen b2 = { b3 where B is Element of b2 : for b1, b2 being Polynomial of b1 holds
( b4 = b3 & b5 in b2 implies len b4 <= len b5 )
}
;

theorem Th17: :: HILBASIS:17
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2 being non empty Subset of (Polynom-Ring b1)
for b3, b4 being Polynomial of b1 holds
( b3 in minlen b2 & b4 in b2 implies ( b3 in b2 & len b3 <= len b4 ) )
proof end;

definition
let c1 be non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr ;
let c2 be Nat;
let c3 be Element of c1;
func monomial c3,c2 -> Polynomial of a1 means :Def5: :: HILBASIS:def 5
for b1 being Nat holds
( ( b1 = a2 implies a4 . b1 = a3 ) & ( b1 <> a2 implies a4 . b1 = 0. a1 ) );
existence
ex b1 being Polynomial of c1 st
for b2 being Nat holds
( ( b2 = c2 implies b1 . b2 = c3 ) & ( b2 <> c2 implies b1 . b2 = 0. c1 ) )
proof end;
uniqueness
for b1, b2 being Polynomial of c1 holds
( ( for b3 being Nat holds
( ( b3 = c2 implies b1 . b3 = c3 ) & ( b3 <> c2 implies b1 . b3 = 0. c1 ) ) ) & ( for b3 being Nat holds
( ( b3 = c2 implies b2 . b3 = c3 ) & ( b3 <> c2 implies b2 . b3 = 0. c1 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def5 defines monomial HILBASIS:def 5 :
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2 being Nat
for b3 being Element of b1
for b4 being Polynomial of b1 holds
( b4 = monomial b3,b2 iff for b5 being Nat holds
( ( b5 = b2 implies b4 . b5 = b3 ) & ( b5 <> b2 implies b4 . b5 = 0. b1 ) ) );

theorem Th18: :: HILBASIS:18
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2 being Nat
for b3 being Element of b1 holds
( ( b3 <> 0. b1 implies len (monomial b3,b2) = b2 + 1 ) & ( b3 = 0. b1 implies len (monomial b3,b2) = 0 ) & len (monomial b3,b2) <= b2 + 1 )
proof end;

theorem Th19: :: HILBASIS:19
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2, b3 being Nat
for b4 being Element of b1
for b5 being Polynomial of b1 holds ((monomial b4,b2) *' b5) . (b3 + b2) = b4 * (b5 . b3)
proof end;

theorem Th20: :: HILBASIS:20
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2, b3 being Nat
for b4 being Element of b1
for b5 being Polynomial of b1 holds (b5 *' (monomial b4,b2)) . (b3 + b2) = (b5 . b3) * b4
proof end;

theorem Th21: :: HILBASIS:21
for b1 being non empty add-associative right_zeroed right_complementable unital distributive doubleLoopStr
for b2, b3 being Polynomial of b1 holds len (b2 *' b3) <= ((len b2) + (len b3)) -' 1
proof end;

theorem Th22: :: HILBASIS:22
for b1, b2 being non empty doubleLoopStr
for b3 being Ideal of b1
for b4 being Function of b1,b2 holds
( b4 is RingIsomorphism implies b4 .: b3 is Ideal of b2 )
proof end;

theorem Th23: :: HILBASIS:23
for b1, b2 being non empty add-associative right_zeroed right_complementable doubleLoopStr
for b3 being Function of b1,b2 holds
( b3 is RingHomomorphism implies b3 . (0. b1) = 0. b2 )
proof end;

theorem Th24: :: HILBASIS:24
for b1, b2 being non empty add-associative right_zeroed right_complementable doubleLoopStr
for b3 being non empty Subset of b1
for b4 being non empty Subset of b2
for b5 being Function of b1,b2
for b6 being LinearCombination of b3
for b7 being LinearCombination of b4
for b8 being FinSequence of [:the carrier of b1,the carrier of b1,the carrier of b1:] holds
( b5 is RingHomomorphism & len b6 = len b7 & b8 represents b6 & ( for b9 being set holds
( b9 in dom b7 implies b7 . b9 = ((b5 . ((b8 /. b9) `1 )) * (b5 . ((b8 /. b9) `2 ))) * (b5 . ((b8 /. b9) `3 )) ) ) implies b5 . (Sum b6) = Sum b7 )
proof end;

theorem Th25: :: HILBASIS:25
for b1, b2 being non empty doubleLoopStr
for b3 being Function of b1,b2 holds
not ( b3 is RingIsomorphism & ( for b4 being Function of b2,b1 holds
not ( b4 is RingIsomorphism & b4 = b3 " ) ) )
proof end;

theorem Th26: :: HILBASIS:26
for b1, b2 being non empty Abelian add-associative right_zeroed right_complementable unital associative distributive doubleLoopStr
for b3 being non empty Subset of b1
for b4 being Function of b1,b2 holds
( b4 is RingIsomorphism implies b4 .: (b3 -Ideal ) = (b4 .: b3) -Ideal )
proof end;

theorem Th27: :: HILBASIS:27
for b1, b2 being non empty Abelian add-associative right_zeroed right_complementable unital associative distributive doubleLoopStr
for b3 being Function of b1,b2 holds
( b3 is RingIsomorphism & b1 is Noetherian implies b2 is Noetherian )
proof end;

theorem Th28: :: HILBASIS:28
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative distributive non trivial doubleLoopStr holds
ex b2 being Function of b1,(Polynom-Ring 0,b1) st b2 is RingIsomorphism
proof end;

theorem Th29: :: HILBASIS:29
for b1 being non empty add-associative right_zeroed right_complementable unital distributive non trivial doubleLoopStr
for b2 being Nat
for b3 being bag of b2
for b4 being Polynomial of b2,b1
for b5 being FinSequence of the carrier of (Polynom-Ring b2,b1) holds
not ( b4 = Sum b5 & ( for b6 being Function of the carrier of (Polynom-Ring b2,b1),the carrier of b1 holds
not ( ( for b7 being Polynomial of b2,b1 holds b6 . b7 = b7 . b3 ) & b4 . b3 = Sum (b6 * b5) ) ) )
proof end;

definition
let c1 be non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr ;
let c2 be Nat;
func upm c2,c1 -> Function of (Polynom-Ring (Polynom-Ring a2,a1)),(Polynom-Ring (a2 + 1),a1) means :Def6: :: HILBASIS:def 6
for b1 being Polynomial of (Polynom-Ring a2,a1)
for b2 being Polynomial of a2,a1
for b3 being Polynomial of (a2 + 1),a1
for b4 being bag of a2 + 1 holds
( b3 = a3 . b1 & b2 = b1 . (b4 . a2) implies b3 . b4 = b2 . (b4 | a2) );
existence
ex b1 being Function of (Polynom-Ring (Polynom-Ring c2,c1)),(Polynom-Ring (c2 + 1),c1) st
for b2 being Polynomial of (Polynom-Ring c2,c1)
for b3 being Polynomial of c2,c1
for b4 being Polynomial of (c2 + 1),c1
for b5 being bag of c2 + 1 holds
( b4 = b1 . b2 & b3 = b2 . (b5 . c2) implies b4 . b5 = b3 . (b5 | c2) )
proof end;
uniqueness
for b1, b2 being Function of (Polynom-Ring (Polynom-Ring c2,c1)),(Polynom-Ring (c2 + 1),c1) holds
( ( for b3 being Polynomial of (Polynom-Ring c2,c1)
for b4 being Polynomial of c2,c1
for b5 being Polynomial of (c2 + 1),c1
for b6 being bag of c2 + 1 holds
( b5 = b1 . b3 & b4 = b3 . (b6 . c2) implies b5 . b6 = b4 . (b6 | c2) ) ) & ( for b3 being Polynomial of (Polynom-Ring c2,c1)
for b4 being Polynomial of c2,c1
for b5 being Polynomial of (c2 + 1),c1
for b6 being bag of c2 + 1 holds
( b5 = b2 . b3 & b4 = b3 . (b6 . c2) implies b5 . b6 = b4 . (b6 | c2) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def6 defines upm HILBASIS:def 6 :
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr
for b2 being Nat
for b3 being Function of (Polynom-Ring (Polynom-Ring b2,b1)),(Polynom-Ring (b2 + 1),b1) holds
( b3 = upm b2,b1 iff for b4 being Polynomial of (Polynom-Ring b2,b1)
for b5 being Polynomial of b2,b1
for b6 being Polynomial of (b2 + 1),b1
for b7 being bag of b2 + 1 holds
( b6 = b3 . b4 & b5 = b4 . (b7 . b2) implies b6 . b7 = b5 . (b7 | b2) ) );

registration
let c1 be non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr ;
let c2 be Nat;
cluster upm a2,a1 -> additive ;
coherence
upm c2,c1 is additive
proof end;
cluster upm a2,a1 -> multiplicative ;
coherence
upm c2,c1 is multiplicative
proof end;
cluster upm a2,a1 -> unity-preserving ;
coherence
upm c2,c1 is unity-preserving
proof end;
cluster upm a2,a1 -> one-to-one ;
coherence
upm c2,c1 is one-to-one
proof end;
end;

definition
let c1 be non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr ;
let c2 be Nat;
func mpu c2,c1 -> Function of (Polynom-Ring (a2 + 1),a1),(Polynom-Ring (Polynom-Ring a2,a1)) means :Def7: :: HILBASIS:def 7
for b1 being Polynomial of (a2 + 1),a1
for b2 being Polynomial of a2,a1
for b3 being Polynomial of (Polynom-Ring a2,a1)
for b4 being Nat
for b5 being bag of a2 holds
( b3 = a3 . b1 & b2 = b3 . b4 implies b2 . b5 = b1 . (b5 bag_extend b4) );
existence
ex b1 being Function of (Polynom-Ring (c2 + 1),c1),(Polynom-Ring (Polynom-Ring c2,c1)) st
for b2 being Polynomial of (c2 + 1),c1
for b3 being Polynomial of c2,c1
for b4 being Polynomial of (Polynom-Ring c2,c1)
for b5 being Nat
for b6 being bag of c2 holds
( b4 = b1 . b2 & b3 = b4 . b5 implies b3 . b6 = b2 . (b6 bag_extend b5) )
proof end;
uniqueness
for b1, b2 being Function of (Polynom-Ring (c2 + 1),c1),(Polynom-Ring (Polynom-Ring c2,c1)) holds
( ( for b3 being Polynomial of (c2 + 1),c1
for b4 being Polynomial of c2,c1
for b5 being Polynomial of (Polynom-Ring c2,c1)
for b6 being Nat
for b7 being bag of c2 holds
( b5 = b1 . b3 & b4 = b5 . b6 implies b4 . b7 = b3 . (b7 bag_extend b6) ) ) & ( for b3 being Polynomial of (c2 + 1),c1
for b4 being Polynomial of c2,c1
for b5 being Polynomial of (Polynom-Ring c2,c1)
for b6 being Nat
for b7 being bag of c2 holds
( b5 = b2 . b3 & b4 = b5 . b6 implies b4 . b7 = b3 . (b7 bag_extend b6) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def7 defines mpu HILBASIS:def 7 :
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr
for b2 being Nat
for b3 being Function of (Polynom-Ring (b2 + 1),b1),(Polynom-Ring (Polynom-Ring b2,b1)) holds
( b3 = mpu b2,b1 iff for b4 being Polynomial of (b2 + 1),b1
for b5 being Polynomial of b2,b1
for b6 being Polynomial of (Polynom-Ring b2,b1)
for b7 being Nat
for b8 being bag of b2 holds
( b6 = b3 . b4 & b5 = b6 . b7 implies b5 . b8 = b4 . (b8 bag_extend b7) ) );

theorem Th30: :: HILBASIS:30
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr
for b2 being Nat
for b3 being Element of (Polynom-Ring (b2 + 1),b1) holds (upm b2,b1) . ((mpu b2,b1) . b3) = b3
proof end;

theorem Th31: :: HILBASIS:31
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr
for b2 being Nat holds
ex b3 being Function of (Polynom-Ring (Polynom-Ring b2,b1)),(Polynom-Ring (b2 + 1),b1) st b3 is RingIsomorphism
proof end;

registration
let c1 be non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive Noetherian doubleLoopStr ;
cluster Polynom-Ring a1 -> Noetherian ;
coherence
Polynom-Ring c1 is Noetherian
proof end;
end;

Lemma36: for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive Noetherian doubleLoopStr holds Polynom-Ring b1 is Noetherian
;

theorem Th32: :: HILBASIS:32
canceled;

theorem Th33: :: HILBASIS:33
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr holds
( b1 is Noetherian implies for b2 being Nat holds Polynom-Ring b2,b1 is Noetherian )
proof end;

theorem Th34: :: HILBASIS:34
for b1 being Field holds b1 is Noetherian
proof end;

theorem Th35: :: HILBASIS:35
for b1 being Field
for b2 being Nat holds Polynom-Ring b2,b1 is Noetherian
proof end;

theorem Th36: :: HILBASIS:36
for b1 being non empty Abelian add-associative right_zeroed right_complementable unital associative commutative distributive non trivial doubleLoopStr
for b2 being infinite Ordinal holds
not Polynom-Ring b2,b1 is Noetherian
proof end;