:: XREAL_0 semantic presentation

definition
let c1 be number ;
attr a1 is real means :Def1: :: XREAL_0:def 1
a1 in REAL ;
end;

:: deftheorem Def1 defines real XREAL_0:def 1 :
for b1 being number holds
( b1 is real iff b1 in REAL );

registration
cluster -infty -> non real ;
coherence
not -infty is real
proof end;
cluster +infty -> non real ;
coherence
not +infty is real
proof end;
end;

registration
cluster natural -> real set ;
coherence
for b1 being number holds
( b1 is natural implies b1 is real )
proof end;
cluster real -> complex set ;
coherence
for b1 being number holds
( b1 is real implies b1 is complex )
proof end;
end;

registration
cluster complex real set ;
existence
ex b1 being number st b1 is real
proof end;
cluster real -> ext-real set ;
coherence
for b1 being number holds
( b1 is real implies b1 is ext-real )
proof end;
end;

Lemma2: for b1 being real number
for b2, b3 being Element of REAL holds
( b1 = [*b2,b3*] implies ( b3 = 0 & b1 = b2 ) )
proof end;

Lemma3: 1 = succ 0
;

registration
let c1 be real number ;
cluster - a1 -> ext-real real ;
coherence
- c1 is real
proof end;
cluster a1 " -> ext-real real ;
coherence
c1 " is real
proof end;
let c2 be real number ;
cluster a1 + a2 -> ext-real real ;
coherence
c1 + c2 is real
proof end;
cluster a1 * a2 -> ext-real real ;
coherence
c1 * c2 is real
proof end;
end;

registration
let c1, c2 be real number ;
cluster a1 - a2 -> ext-real real ;
coherence
c1 - c2 is real
;
cluster a1 / a2 -> ext-real real ;
coherence
c1 / c2 is real
;
end;

Lemma4: for b1, b2 being real number holds
( b1 <= b2 implies ( not ( b1 in REAL+ & b2 in REAL+ & ( for b3, b4 being Element of REAL+ holds
not ( b1 = b3 & b2 = b4 & b3 <=' b4 ) ) ) & not ( b1 in [:{0},REAL+ :] & b2 in [:{0},REAL+ :] & ( for b3, b4 being Element of REAL+ holds
not ( b1 = [0,b3] & b2 = [0,b4] & b4 <=' b3 ) ) ) & ( not ( b1 in REAL+ & b2 in REAL+ ) & not ( b1 in [:{0},REAL+ :] & b2 in [:{0},REAL+ :] ) implies ( b2 in REAL+ & b1 in [:{0},REAL+ :] ) ) ) )
by XXREAL_0:def 5;

Lemma5: for b1, b2 being real number holds
( not ( not ( b1 in REAL+ & b2 in REAL+ & ex b3, b4 being Element of REAL+ st
( b1 = b3 & b2 = b4 & b3 <=' b4 ) ) & not ( b1 in [:{0},REAL+ :] & b2 in [:{0},REAL+ :] & ex b3, b4 being Element of REAL+ st
( b1 = [0,b3] & b2 = [0,b4] & b4 <=' b3 ) ) & not ( b2 in REAL+ & b1 in [:{0},REAL+ :] ) ) implies b1 <= b2 )
proof end;

Lemma6: {} in {{} }
by TARSKI:def 1;

Lemma7: for b1, b2 being real number holds
( b1 <= b2 & b2 <= b1 implies b1 = b2 )
proof end;

Lemma8: for b1, b2, b3 being real number holds
( b1 <= b2 implies b1 + b3 <= b2 + b3 )
proof end;

Lemma9: for b1, b2, b3 being real number holds
( b1 <= b2 & b2 <= b3 implies b1 <= b3 )
proof end;

reconsider c1 = 0 as Element of REAL+ by ARYTM_2:21;

Lemma10: not 0 in [:{0},REAL+ :]
by ARYTM_0:5, ARYTM_2:21, XBOOLE_0:3;

reconsider c2 = 1 as Element of REAL+ by Lemma3, ARYTM_2:21;

c1 <=' c2
by ARYTM_1:6;

then Lemma11: 0 <= 1
by Lemma5;

1 + (- 1) = 0
;

then consider c3, c4, c5, c6 being Element of REAL such that
Lemma12: 1 = [*c3,c4*] and
Lemma13: - 1 = [*c5,c6*] and
Lemma14: 0 = [*(+ c3,c5),(+ c4,c6)*] by XCMPLX_0:def 4;

Lemma15: c3 = 1
by Lemma2, Lemma12;

Lemma16: c5 = - 1
by Lemma2, Lemma13;

Lemma17: + c3,c5 = 0
by Lemma2, Lemma14;

E18: now
assume - 1 in REAL+ ;
then ex b1, b2 being Element of REAL+ st
( c3 = b1 & c5 = b2 & c1 = b1 + b2 ) by Lemma3, Lemma15, Lemma16, Lemma17, ARYTM_0:def 2, ARYTM_2:21;
hence not verum by Lemma15, ARYTM_2:6;
end;

Lemma19: for b1, b2 being real number holds
not ( b1 >= 0 & b2 > 0 & not b1 + b2 > 0 )
proof end;

Lemma20: for b1, b2 being real number holds
not ( b1 <= 0 & b2 < 0 & not b1 + b2 < 0 )
proof end;

reconsider c7 = 0 as Element of REAL+ by ARYTM_2:21;

Lemma21: for b1, b2, b3 being real number holds
( b1 <= b2 & 0 <= b3 implies b1 * b3 <= b2 * b3 )
proof end;

Lemma22: for b1, b2, b3 being real number holds (b1 * b2) * b3 = b1 * (b2 * b3)
;

Lemma23: for b1, b2 being real number holds
not ( b1 * b2 = 0 & not b1 = 0 & not b2 = 0 )
proof end;

Lemma24: for b1, b2 being real number holds
not ( b1 > 0 & b2 > 0 & not b1 * b2 > 0 )
proof end;

Lemma25: for b1, b2 being real number holds
not ( b1 > 0 & b2 < 0 & not b1 * b2 < 0 )
proof end;

Lemma26: for b1, b2 being real number holds
( b1 <= b2 implies - b2 <= - b1 )
proof end;

Lemma27: for b1, b2 being real number holds
( b1 <= 0 & b2 >= 0 implies b1 * b2 <= 0 )
proof end;

registration
cluster complex ext-real positive real set ;
existence
ex b1 being real number st b1 is positive
proof end;
cluster complex ext-real negative real set ;
existence
ex b1 being real number st b1 is negative
proof end;
cluster zero complex ext-real real set ;
existence
ex b1 being real number st b1 is empty
proof end;
end;

registration
let c8, c9 be non negative real number ;
cluster a1 + a2 -> ext-real non negative real ;
coherence
not c8 + c9 is negative
proof end;
end;

registration
let c8, c9 be non positive real number ;
cluster a1 + a2 -> ext-real non positive real ;
coherence
not c8 + c9 is positive
proof end;
end;

registration
let c8 be positive real number ;
let c9 be non negative real number ;
cluster a1 + a2 -> ext-real positive non negative real ;
coherence
c8 + c9 is positive
proof end;
cluster a2 + a1 -> ext-real positive non negative real ;
coherence
c9 + c8 is positive
;
end;

registration
let c8 be negative real number ;
let c9 be non positive real number ;
cluster a1 + a2 -> ext-real non positive negative real ;
coherence
c8 + c9 is negative
proof end;
cluster a2 + a1 -> ext-real non positive negative real ;
coherence
c9 + c8 is negative
;
end;

registration
let c8 be non positive real number ;
cluster - a1 -> ext-real non negative real ;
coherence
not - c8 is negative
proof end;
end;

registration
let c8 be non negative real number ;
cluster - a1 -> ext-real non positive real ;
coherence
not - c8 is positive
proof end;
end;

registration
let c8 be non negative real number ;
let c9 be non positive real number ;
cluster a1 - a2 -> ext-real non negative real ;
coherence
not c8 - c9 is negative
;
cluster a2 - a1 -> ext-real non positive real ;
coherence
not c9 - c8 is positive
;
end;

registration
let c8 be positive real number ;
let c9 be non positive real number ;
cluster a1 - a2 -> ext-real positive non negative real ;
coherence
c8 - c9 is positive
;
cluster a2 - a1 -> ext-real non positive negative real ;
coherence
c9 - c8 is negative
;
end;

registration
let c8 be negative real number ;
let c9 be non negative real number ;
cluster a1 - a2 -> ext-real non positive negative real ;
coherence
c8 - c9 is negative
;
cluster a2 - a1 -> ext-real positive non negative real ;
coherence
c9 - c8 is positive
;
end;

registration
let c8 be non positive real number ;
let c9 be non negative real number ;
cluster a1 * a2 -> ext-real non positive real ;
coherence
not c8 * c9 is positive
proof end;
cluster a2 * a1 -> ext-real non positive real ;
coherence
not c9 * c8 is positive
;
end;

registration
let c8, c9 be non positive real number ;
cluster a1 * a2 -> ext-real non negative real ;
coherence
not c8 * c9 is negative
proof end;
end;

registration
let c8, c9 be non negative real number ;
cluster a1 * a2 -> ext-real non negative real ;
coherence
not c8 * c9 is negative
proof end;
end;

Lemma28: for b1 being real number holds
( b1 " = 0 implies b1 = 0 )
proof end;

registration
let c8 be non positive real number ;
cluster a1 " -> ext-real non positive real ;
coherence
not c8 " is positive
proof end;
end;

registration
let c8 be non negative real number ;
cluster a1 " -> ext-real non negative real ;
coherence
not c8 " is negative
proof end;
end;

registration
let c8 be non negative real number ;
let c9 be non positive real number ;
cluster a1 / a2 -> ext-real non positive real ;
coherence
not c8 / c9 is positive
;
cluster a2 / a1 -> ext-real non positive real ;
coherence
not c9 / c8 is positive
;
end;

registration
let c8, c9 be non negative real number ;
cluster a1 / a2 -> ext-real non negative real ;
coherence
not c8 / c9 is negative
;
end;

registration
let c8, c9 be non positive real number ;
cluster a1 / a2 -> ext-real non negative real ;
coherence
not c8 / c9 is negative
;
end;

registration
let c8, c9 be real number ;
cluster min a1,a2 -> complex real ;
coherence
min c8,c9 is real
by XXREAL_0:15;
cluster max a1,a2 -> complex real ;
coherence
max c8,c9 is real
by XXREAL_0:16;
end;