begin
:: deftheorem Def1 defines <= XXREAL_3:def 1 :
for x, y being ext-real number holds
( ( x in REAL & y in REAL implies ( x <= y iff ex p, q being Element of REAL st
( p = x & q = y & p <= q ) ) ) & ( ( not x in REAL or not y in REAL ) implies ( x <= y iff ( x = -infty or y = +infty ) ) ) );
theorem Th1:
theorem Th2:
theorem Th3:
begin
definition
let x,
y be
ext-real number ;
func x + y -> ext-real number means :
Def2:
ex
a,
b being
complex number st
(
x = a &
y = b &
it = a + b )
if (
x is
real &
y is
real )
it = +infty if ( (
x = +infty &
y <> -infty ) or (
y = +infty &
x <> -infty ) )
it = -infty if ( (
x = -infty &
y <> +infty ) or (
y = -infty &
x <> +infty ) )
otherwise it = 0 ;
existence
( ( x is real & y is real implies ex b1 being ext-real number ex a, b being complex number st
( x = a & y = b & b1 = a + b ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies ex b1 being ext-real number st b1 = +infty ) & ( ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ex b1 being ext-real number st b1 = -infty ) & ( ( x is real & y is real ) or ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) or ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) or ex b1 being ext-real number st b1 = 0 ) )
uniqueness
for b1, b2 being ext-real number holds
( ( x is real & y is real & ex a, b being complex number st
( x = a & y = b & b1 = a + b ) & ex a, b being complex number st
( x = a & y = b & b2 = a + b ) implies b1 = b2 ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) & b1 = +infty & b2 = +infty implies b1 = b2 ) & ( ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) & b1 = -infty & b2 = -infty implies b1 = b2 ) & ( ( x is real & y is real ) or ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) or ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) or not b1 = 0 or not b2 = 0 or b1 = b2 ) )
;
consistency
for b1 being ext-real number holds
( ( x is real & y is real & ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies ( ex a, b being complex number st
( x = a & y = b & b1 = a + b ) iff b1 = +infty ) ) & ( x is real & y is real & ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( ex a, b being complex number st
( x = a & y = b & b1 = a + b ) iff b1 = -infty ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) & ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( b1 = +infty iff b1 = -infty ) ) )
;
commutativity
for b1, x, y being ext-real number st ( x is real & y is real implies ex a, b being complex number st
( x = a & y = b & b1 = a + b ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies b1 = +infty ) & ( ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies b1 = -infty ) & ( ( x is real & y is real ) or ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) or ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) or b1 = 0 ) holds
( ( y is real & x is real implies ex a, b being complex number st
( y = a & x = b & b1 = a + b ) ) & ( ( ( y = +infty & x <> -infty ) or ( x = +infty & y <> -infty ) ) implies b1 = +infty ) & ( ( ( y = -infty & x <> +infty ) or ( x = -infty & y <> +infty ) ) implies b1 = -infty ) & ( ( y is real & x is real ) or ( y = +infty & x <> -infty ) or ( x = +infty & y <> -infty ) or ( y = -infty & x <> +infty ) or ( x = -infty & y <> +infty ) or b1 = 0 ) )
;
end;
:: deftheorem Def2 defines + XXREAL_3:def 2 :
for x, y, b3 being ext-real number holds
( ( x is real & y is real implies ( b3 = x + y iff ex a, b being complex number st
( x = a & y = b & b3 = a + b ) ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies ( b3 = x + y iff b3 = +infty ) ) & ( ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( b3 = x + y iff b3 = -infty ) ) & ( ( x is real & y is real ) or ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) or ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) or ( b3 = x + y iff b3 = 0 ) ) );
:: deftheorem Def3 defines - XXREAL_3:def 3 :
for x, b2 being ext-real number holds
( ( x is real implies ( b2 = - x iff ex a being complex number st
( x = a & b2 = - a ) ) ) & ( x = +infty implies ( b2 = - x iff b2 = -infty ) ) & ( not x is real & not x = +infty implies ( b2 = - x iff b2 = +infty ) ) );
:: deftheorem defines - XXREAL_3:def 4 :
for x, y being ext-real number holds x - y = x + (- y);
theorem Th4:
theorem Th5:
theorem Th6:
Lm1:
+infty + +infty = +infty
by Def2;
Lm2:
-infty + -infty = -infty
by Def2;
theorem Th7:
theorem
Lm3:
for x being ext-real number holds
( - x in REAL iff x in REAL )
Lm4: - (+infty + -infty) =
+infty - +infty
by Def3
.=
(- -infty) - +infty
by Def3
;
Lm5:
- +infty = -infty
by Def3;
Lm6:
for x being ext-real number st x in REAL holds
- (x + +infty) = (- +infty) + (- x)
Lm7:
for x being ext-real number st x in REAL holds
- (x + -infty) = (- -infty) + (- x)
theorem Th9:
theorem Th10:
Lm8:
for x, y being ext-real number holds
( not x + y = +infty or x = +infty or y = +infty )
Lm9:
for x, y being ext-real number holds
( not x + y = -infty or x = -infty or y = -infty )
theorem Th11:
theorem
theorem Th13:
theorem Th14:
theorem Th15:
theorem
theorem
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem
Lm10:
- +infty = -infty
by Def3;
Lm11:
for x being ext-real number st x in REAL holds
- (x + -infty) = (- -infty) + (- x)
theorem
canceled;
theorem Th26:
theorem
theorem Th28:
theorem
theorem Th30:
theorem Th31:
theorem
theorem Th33:
theorem Th34:
theorem
begin
Lm12:
for x, y, z being ext-real number st x <= y holds
x + z <= y + z
Lm13:
for x, y being ext-real number st x >= 0 & y > 0 holds
x + y > 0
Lm14:
for x, y being ext-real number st x <= 0 & y < 0 holds
x + y < 0
Lm15:
for x, y being ext-real number st x <= y holds
- y <= - x
theorem
theorem
canceled;
theorem Th38:
theorem Th39:
theorem Th40:
theorem
canceled;
theorem Th42:
theorem
theorem
theorem
theorem
canceled;
theorem Th47:
theorem
theorem Th49:
theorem Th50:
theorem
theorem
canceled;
theorem
theorem
canceled;
theorem Th55:
theorem
theorem Th57:
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
begin
definition
let x,
y be
ext-real number ;
func x * y -> ext-real number means :
Def5:
ex
a,
b being
complex number st
(
x = a &
y = b &
it = a * b )
if (
x is
real &
y is
real )
it = +infty if ( ( not
x is
real or not
y is
real ) & ( (
x is
positive &
y is
positive ) or (
x is
negative &
y is
negative ) ) )
it = -infty if ( ( not
x is
real or not
y is
real ) & ( (
x is
positive &
y is
negative ) or (
x is
negative &
y is
positive ) ) )
otherwise it = 0 ;
existence
( ( x is real & y is real implies ex b1 being ext-real number ex a, b being complex number st
( x = a & y = b & b1 = a * b ) ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is positive ) or ( x is negative & y is negative ) ) implies ex b1 being ext-real number st b1 = +infty ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is negative ) or ( x is negative & y is positive ) ) implies ex b1 being ext-real number st b1 = -infty ) & ( ( not x is real or not y is real ) & ( ( x is real & y is real ) or ( not ( x is positive & y is positive ) & not ( x is negative & y is negative ) ) ) & ( ( x is real & y is real ) or ( not ( x is positive & y is negative ) & not ( x is negative & y is positive ) ) ) implies ex b1 being ext-real number st b1 = 0 ) )
uniqueness
for b1, b2 being ext-real number holds
( ( x is real & y is real & ex a, b being complex number st
( x = a & y = b & b1 = a * b ) & ex a, b being complex number st
( x = a & y = b & b2 = a * b ) implies b1 = b2 ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is positive ) or ( x is negative & y is negative ) ) & b1 = +infty & b2 = +infty implies b1 = b2 ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is negative ) or ( x is negative & y is positive ) ) & b1 = -infty & b2 = -infty implies b1 = b2 ) & ( ( not x is real or not y is real ) & ( ( x is real & y is real ) or ( not ( x is positive & y is positive ) & not ( x is negative & y is negative ) ) ) & ( ( x is real & y is real ) or ( not ( x is positive & y is negative ) & not ( x is negative & y is positive ) ) ) & b1 = 0 & b2 = 0 implies b1 = b2 ) )
;
consistency
for b1 being ext-real number holds
( ( x is real & y is real & ( not x is real or not y is real ) & ( ( x is positive & y is positive ) or ( x is negative & y is negative ) ) implies ( ex a, b being complex number st
( x = a & y = b & b1 = a * b ) iff b1 = +infty ) ) & ( x is real & y is real & ( not x is real or not y is real ) & ( ( x is positive & y is negative ) or ( x is negative & y is positive ) ) implies ( ex a, b being complex number st
( x = a & y = b & b1 = a * b ) iff b1 = -infty ) ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is positive ) or ( x is negative & y is negative ) ) & ( not x is real or not y is real ) & ( ( x is positive & y is negative ) or ( x is negative & y is positive ) ) implies ( b1 = +infty iff b1 = -infty ) ) )
;
commutativity
for b1, x, y being ext-real number st ( x is real & y is real implies ex a, b being complex number st
( x = a & y = b & b1 = a * b ) ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is positive ) or ( x is negative & y is negative ) ) implies b1 = +infty ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is negative ) or ( x is negative & y is positive ) ) implies b1 = -infty ) & ( ( not x is real or not y is real ) & ( ( x is real & y is real ) or ( not ( x is positive & y is positive ) & not ( x is negative & y is negative ) ) ) & ( ( x is real & y is real ) or ( not ( x is positive & y is negative ) & not ( x is negative & y is positive ) ) ) implies b1 = 0 ) holds
( ( y is real & x is real implies ex a, b being complex number st
( y = a & x = b & b1 = a * b ) ) & ( ( not y is real or not x is real ) & ( ( y is positive & x is positive ) or ( y is negative & x is negative ) ) implies b1 = +infty ) & ( ( not y is real or not x is real ) & ( ( y is positive & x is negative ) or ( y is negative & x is positive ) ) implies b1 = -infty ) & ( ( not y is real or not x is real ) & ( ( y is real & x is real ) or ( not ( y is positive & x is positive ) & not ( y is negative & x is negative ) ) ) & ( ( y is real & x is real ) or ( not ( y is positive & x is negative ) & not ( y is negative & x is positive ) ) ) implies b1 = 0 ) )
;
end;
:: deftheorem Def5 defines * XXREAL_3:def 5 :
for x, y, b3 being ext-real number holds
( ( x is real & y is real implies ( b3 = x * y iff ex a, b being complex number st
( x = a & y = b & b3 = a * b ) ) ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is positive ) or ( x is negative & y is negative ) ) implies ( b3 = x * y iff b3 = +infty ) ) & ( ( not x is real or not y is real ) & ( ( x is positive & y is negative ) or ( x is negative & y is positive ) ) implies ( b3 = x * y iff b3 = -infty ) ) & ( ( not x is real or not y is real ) & ( ( x is real & y is real ) or ( not ( x is positive & y is positive ) & not ( x is negative & y is negative ) ) ) & ( ( x is real & y is real ) or ( not ( x is positive & y is negative ) & not ( x is negative & y is positive ) ) ) implies ( b3 = x * y iff b3 = 0 ) ) );
:: deftheorem Def6 defines " XXREAL_3:def 6 :
for x, b2 being ext-real number holds
( ( x is real implies ( b2 = x " iff ex a being complex number st
( x = a & b2 = a " ) ) ) & ( not x is real implies ( b2 = x " iff b2 = 0 ) ) );
:: deftheorem defines / XXREAL_3:def 7 :
for x, y being ext-real number holds x / y = x * (y ");
Lm16:
for x being ext-real number holds x * 0 = 0
Lm17:
for x, y being ext-real number st not x is real & x * y = 0 holds
y = 0
Lm18:
for x, y, z being ext-real number st x = 0 holds
x * (y * z) = (x * y) * z
Lm19:
for y, x, z being ext-real number st y = 0 holds
x * (y * z) = (x * y) * z
Lm20:
for y, x, z being ext-real number st not y is real holds
x * (y * z) = (x * y) * z
Lm21:
for x, y, z being ext-real number st not x is real holds
x * (y * z) = (x * y) * z
theorem Th77:
theorem
theorem
theorem Th80:
theorem Th81:
Lm22:
for x, y being ext-real number holds
( not x * y in REAL or ( x in REAL & y in REAL ) or x * y = 0 )
theorem Th82:
theorem Th83:
theorem
theorem
theorem
theorem
theorem
theorem Th89:
theorem
theorem
theorem Th92:
theorem Th93:
theorem Th94:
theorem Th95:
theorem Th96:
theorem Th97:
theorem
theorem
theorem
theorem
begin
theorem Th102:
theorem Th103:
theorem Th104:
theorem Th105:
Lm23:
for x, y being ext-real number holds x * (y + y) = (x * y) + (x * y)
Lm24:
for x, z being ext-real number holds x * (0 + z) = (x * 0) + (x * z)
Lm25:
for y, z being ext-real number holds 0 * (y + z) = (0 * y) + (0 * z)
;
Lm26:
for x, y, z being ext-real number st x is real & y is real holds
x * (y + z) = (x * y) + (x * z)
Lm27:
for x, y, z being ext-real number st x is real & not y is real holds
x * (y + z) = (x * y) + (x * z)
theorem Th106:
theorem Th107:
theorem
canceled;
theorem
theorem
theorem
theorem
theorem Th113:
theorem Th114:
theorem
theorem
theorem