Database
REAL AND COMPLEX NUMBERS
Construction and axiomatization of real and complex numbers
Real and complex number postulates restated as axioms
ax-pre-ltadd
Metamath Proof Explorer
Description: Ordering property of addition on reals. Axiom 20 of 22 for real and
complex numbers, justified by Theorem axpre-ltadd . Normally new proofs
would use axltadd . (New usage is discouraged.) (Contributed by NM , 13-Oct-2005)
Ref
Expression
Assertion
ax-pre-ltadd
⊢ A ∈ ℝ ∧ B ∈ ℝ ∧ C ∈ ℝ → A < ℝ B → C + A < ℝ C + B
Detailed syntax breakdown
Step
Hyp
Ref
Expression
0
cA
class A
1
cr
class ℝ
2
0 1
wcel
wff A ∈ ℝ
3
cB
class B
4
3 1
wcel
wff B ∈ ℝ
5
cC
class C
6
5 1
wcel
wff C ∈ ℝ
7
2 4 6
w3a
wff A ∈ ℝ ∧ B ∈ ℝ ∧ C ∈ ℝ
8
cltrr
class < ℝ
9
0 3 8
wbr
wff A < ℝ B
10
caddc
class +
11
5 0 10
co
class C + A
12
5 3 10
co
class C + B
13
11 12 8
wbr
wff C + A < ℝ C + B
14
9 13
wi
wff A < ℝ B → C + A < ℝ C + B
15
7 14
wi
wff A ∈ ℝ ∧ B ∈ ℝ ∧ C ∈ ℝ → A < ℝ B → C + A < ℝ C + B