Lm1:
0 = 0 + (0 * <i>)
;
Lm2:
now for V being non empty ModuleStr over F_Complex
for f being Functional of V
for v1, w being Vector of V holds (FormFunctional (f,(0Functional V))) . (v1,w) = 0. F_Complex
let V be non
empty ModuleStr over
F_Complex ;
for f being Functional of V
for v1, w being Vector of V holds (FormFunctional (f,(0Functional V))) . (v1,w) = 0. F_Complexlet f be
Functional of
V;
for v1, w being Vector of V holds (FormFunctional (f,(0Functional V))) . (v1,w) = 0. F_Complexset 0F =
0Functional V;
let v1,
w be
Vector of
V;
(FormFunctional (f,(0Functional V))) . (v1,w) = 0. F_Complexthus (FormFunctional (f,(0Functional V))) . (
v1,
w) =
(f . v1) * ((0Functional V) . w)
by BILINEAR:def 10
.=
(f . v1) * (0. F_Complex)
by HAHNBAN1:14
.=
0. F_Complex
;
verum
end;
Lm3:
for V being non empty ModuleStr over F_Complex
for f being Functional of V holds FormFunctional (f,(0Functional V)) is hermitan
definition
let V,
W be non
empty ModuleStr over
F_Complex ;
let f be
Form of
V,
W;
existence
ex b1 being Form of V,W st
for v being Vector of V
for w being Vector of W holds b1 . (v,w) = (f . (v,w)) *'
uniqueness
for b1, b2 being Form of V,W st ( for v being Vector of V
for w being Vector of W holds b1 . (v,w) = (f . (v,w)) *' ) & ( for v being Vector of V
for w being Vector of W holds b2 . (v,w) = (f . (v,w)) *' ) holds
b1 = b2
end;
theorem Th37:
for
V,
W being non
empty right_complementable add-associative right_zeroed vector-distributive scalar-distributive scalar-associative scalar-unital ModuleStr over
F_Complex for
v,
u being
Vector of
V for
w,
t being
Vector of
W for
a,
b being
Element of
F_Complex for
f being
sesquilinear-Form of
V,
W holds
f . (
(v + (a * u)),
(w + (b * t)))
= ((f . (v,w)) + ((b *') * (f . (v,t)))) + ((a * (f . (u,w))) + (a * ((b *') * (f . (u,t)))))
theorem Th38:
for
V,
W being
VectSp of
F_Complex for
v,
u being
Vector of
V for
w,
t being
Vector of
W for
a,
b being
Element of
F_Complex for
f being
sesquilinear-Form of
V,
W holds
f . (
(v - (a * u)),
(w - (b * t)))
= ((f . (v,w)) - ((b *') * (f . (v,t)))) - ((a * (f . (u,w))) - (a * ((b *') * (f . (u,t)))))
theorem Th42:
for
V being
VectSp of
F_Complex for
v,
w being
Vector of
V for
f being
sesquilinear-Form of
V,
V for
r being
Real for
a being
Element of
F_Complex st
|.a.| = 1 holds
f . (
(v - (([**r,0**] * a) * w)),
(v - (([**r,0**] * a) * w)))
= (((f . (v,v)) - ([**r,0**] * (a * (f . (w,v))))) - ([**r,0**] * ((a *') * (f . (v,w))))) + ([**(r ^2),0**] * (f . (w,w)))
theorem Th43:
for
V being
VectSp of
F_Complex for
v,
w being
Vector of
V for
f being
diagReR+0valued hermitan-Form of
V for
r being
Real for
a being
Element of
F_Complex st
|.a.| = 1 &
Re (a * (f . (w,v))) = |.(f . (w,v)).| holds
(
Re (f . ((v - (([**r,0**] * a) * w)),(v - (([**r,0**] * a) * w)))) = ((signnorm (f,v)) - ((2 * |.(f . (w,v)).|) * r)) + ((signnorm (f,w)) * (r ^2)) &
0 <= ((signnorm (f,v)) - ((2 * |.(f . (w,v)).|) * r)) + ((signnorm (f,w)) * (r ^2)) )
definition
let V,
W be
VectSp of
F_Complex ;
let f be
sesquilinear-Form of
V,
W;
existence
ex b1 being sesquilinear-Form of (VectQuot (V,(LKer f))),(VectQuot (W,(RKer (f *')))) st
for A being Vector of (VectQuot (V,(LKer f)))
for B being Vector of (VectQuot (W,(RKer (f *'))))
for v being Vector of V
for w being Vector of W st A = v + (LKer f) & B = w + (RKer (f *')) holds
b1 . (A,B) = f . (v,w)
uniqueness
for b1, b2 being sesquilinear-Form of (VectQuot (V,(LKer f))),(VectQuot (W,(RKer (f *')))) st ( for A being Vector of (VectQuot (V,(LKer f)))
for B being Vector of (VectQuot (W,(RKer (f *'))))
for v being Vector of V
for w being Vector of W st A = v + (LKer f) & B = w + (RKer (f *')) holds
b1 . (A,B) = f . (v,w) ) & ( for A being Vector of (VectQuot (V,(LKer f)))
for B being Vector of (VectQuot (W,(RKer (f *'))))
for v being Vector of V
for w being Vector of W st A = v + (LKer f) & B = w + (RKer (f *')) holds
b2 . (A,B) = f . (v,w) ) holds
b1 = b2
end;