HomeHome Intuitionistic Logic Explorer
Theorem List (p. 71 of 102)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 7001-7100   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Axiomax-pre-mulgt0 7001 The product of two positive reals is positive. Axiom for real and complex numbers, justified by theorem axpre-mulgt0 6961. (Contributed by NM, 13-Oct-2005.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → ((0 < 𝐴 ∧ 0 < 𝐵) → 0 < (𝐴 · 𝐵)))
 
Axiomax-pre-mulext 7002 Strong extensionality of multiplication (expressed in terms of <). Axiom for real and complex numbers, justified by theorem axpre-mulext 6962

(Contributed by Jim Kingdon, 18-Feb-2020.)

((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → ((𝐴 · 𝐶) < (𝐵 · 𝐶) → (𝐴 < 𝐵𝐵 < 𝐴)))
 
Axiomax-arch 7003* Archimedean axiom. Definition 3.1(2) of [Geuvers], p. 9. Axiom for real and complex numbers, justified by theorem axarch 6965.

This axiom should not be used directly; instead use arch 8178 (which is the same, but stated in terms of and <). (Contributed by Jim Kingdon, 2-May-2020.) (New usage is discouraged.)

(𝐴 ∈ ℝ → ∃𝑛 {𝑥 ∣ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)}𝐴 < 𝑛)
 
Axiomax-caucvg 7004* Completeness. Axiom for real and complex numbers, justified by theorem axcaucvg 6974.

A Cauchy sequence (as defined here, which has a rate convergence built in) of real numbers converges to a real number. Specifically on rate of convergence, all terms after the nth term must be within 1 / 𝑛 of the nth term.

This axiom should not be used directly; instead use caucvgre 9580 (which is the same, but stated in terms of the and 1 / 𝑛 notations). (Contributed by Jim Kingdon, 19-Jul-2021.) (New usage is discouraged.)

𝑁 = {𝑥 ∣ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)}    &   (𝜑𝐹:𝑁⟶ℝ)    &   (𝜑 → ∀𝑛𝑁𝑘𝑁 (𝑛 < 𝑘 → ((𝐹𝑛) < ((𝐹𝑘) + (𝑟 ∈ ℝ (𝑛 · 𝑟) = 1)) ∧ (𝐹𝑘) < ((𝐹𝑛) + (𝑟 ∈ ℝ (𝑛 · 𝑟) = 1)))))       (𝜑 → ∃𝑦 ∈ ℝ ∀𝑥 ∈ ℝ (0 < 𝑥 → ∃𝑗𝑁𝑘𝑁 (𝑗 < 𝑘 → ((𝐹𝑘) < (𝑦 + 𝑥) ∧ 𝑦 < ((𝐹𝑘) + 𝑥)))))
 
3.2  Derive the basic properties from the field axioms
 
3.2.1  Some deductions from the field axioms for complex numbers
 
Theoremcnex 7005 Alias for ax-cnex 6975. (Contributed by Mario Carneiro, 17-Nov-2014.)
ℂ ∈ V
 
Theoremaddcl 7006 Alias for ax-addcl 6980, for naming consistency with addcli 7031. Use this theorem instead of ax-addcl 6980 or axaddcl 6940. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴 + 𝐵) ∈ ℂ)
 
Theoremreaddcl 7007 Alias for ax-addrcl 6981, for naming consistency with readdcli 7040. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 + 𝐵) ∈ ℝ)
 
Theoremmulcl 7008 Alias for ax-mulcl 6982, for naming consistency with mulcli 7032. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴 · 𝐵) ∈ ℂ)
 
Theoremremulcl 7009 Alias for ax-mulrcl 6983, for naming consistency with remulcli 7041. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 · 𝐵) ∈ ℝ)
 
Theoremmulcom 7010 Alias for ax-mulcom 6985, for naming consistency with mulcomi 7033. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴 · 𝐵) = (𝐵 · 𝐴))
 
Theoremaddass 7011 Alias for ax-addass 6986, for naming consistency with addassi 7035. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → ((𝐴 + 𝐵) + 𝐶) = (𝐴 + (𝐵 + 𝐶)))
 
Theoremmulass 7012 Alias for ax-mulass 6987, for naming consistency with mulassi 7036. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → ((𝐴 · 𝐵) · 𝐶) = (𝐴 · (𝐵 · 𝐶)))
 
Theoremadddi 7013 Alias for ax-distr 6988, for naming consistency with adddii 7037. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → (𝐴 · (𝐵 + 𝐶)) = ((𝐴 · 𝐵) + (𝐴 · 𝐶)))
 
Theoremrecn 7014 A real number is a complex number. (Contributed by NM, 10-Aug-1999.)
(𝐴 ∈ ℝ → 𝐴 ∈ ℂ)
 
Theoremreex 7015 The real numbers form a set. (Contributed by Mario Carneiro, 17-Nov-2014.)
ℝ ∈ V
 
Theoremreelprrecn 7016 Reals are a subset of the pair of real and complex numbers (common case). (Contributed by David A. Wheeler, 8-Dec-2018.)
ℝ ∈ {ℝ, ℂ}
 
Theoremcnelprrecn 7017 Complex numbers are a subset of the pair of real and complex numbers (common case). (Contributed by David A. Wheeler, 8-Dec-2018.)
ℂ ∈ {ℝ, ℂ}
 
Theoremadddir 7018 Distributive law for complex numbers (right-distributivity). (Contributed by NM, 10-Oct-2004.)
((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → ((𝐴 + 𝐵) · 𝐶) = ((𝐴 · 𝐶) + (𝐵 · 𝐶)))
 
Theorem0cn 7019 0 is a complex number. (Contributed by NM, 19-Feb-2005.)
0 ∈ ℂ
 
Theorem0cnd 7020 0 is a complex number, deductive form. (Contributed by David A. Wheeler, 8-Dec-2018.)
(𝜑 → 0 ∈ ℂ)
 
Theoremc0ex 7021 0 is a set (common case). (Contributed by David A. Wheeler, 7-Jul-2016.)
0 ∈ V
 
Theorem1ex 7022 1 is a set. Common special case. (Contributed by David A. Wheeler, 7-Jul-2016.)
1 ∈ V
 
Theoremcnre 7023* Alias for ax-cnre 6995, for naming consistency. (Contributed by NM, 3-Jan-2013.)
(𝐴 ∈ ℂ → ∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)))
 
Theoremmulid1 7024 1 is an identity element for multiplication. Based on ideas by Eric Schmidt. (Contributed by Scott Fenton, 3-Jan-2013.)
(𝐴 ∈ ℂ → (𝐴 · 1) = 𝐴)
 
Theoremmulid2 7025 Identity law for multiplication. Note: see mulid1 7024 for commuted version. (Contributed by NM, 8-Oct-1999.)
(𝐴 ∈ ℂ → (1 · 𝐴) = 𝐴)
 
Theorem1re 7026 1 is a real number. (Contributed by Jim Kingdon, 13-Jan-2020.)
1 ∈ ℝ
 
Theorem0re 7027 0 is a real number. (Contributed by Eric Schmidt, 21-May-2007.) (Revised by Scott Fenton, 3-Jan-2013.)
0 ∈ ℝ
 
Theorem0red 7028 0 is a real number, deductive form. (Contributed by David A. Wheeler, 6-Dec-2018.)
(𝜑 → 0 ∈ ℝ)
 
Theoremmulid1i 7029 Identity law for multiplication. (Contributed by NM, 14-Feb-1995.)
𝐴 ∈ ℂ       (𝐴 · 1) = 𝐴
 
Theoremmulid2i 7030 Identity law for multiplication. (Contributed by NM, 14-Feb-1995.)
𝐴 ∈ ℂ       (1 · 𝐴) = 𝐴
 
Theoremaddcli 7031 Closure law for addition. (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ       (𝐴 + 𝐵) ∈ ℂ
 
Theoremmulcli 7032 Closure law for multiplication. (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ       (𝐴 · 𝐵) ∈ ℂ
 
Theoremmulcomi 7033 Commutative law for multiplication. (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ       (𝐴 · 𝐵) = (𝐵 · 𝐴)
 
Theoremmulcomli 7034 Commutative law for multiplication. (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ    &   (𝐴 · 𝐵) = 𝐶       (𝐵 · 𝐴) = 𝐶
 
Theoremaddassi 7035 Associative law for addition. (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ    &   𝐶 ∈ ℂ       ((𝐴 + 𝐵) + 𝐶) = (𝐴 + (𝐵 + 𝐶))
 
Theoremmulassi 7036 Associative law for multiplication. (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ    &   𝐶 ∈ ℂ       ((𝐴 · 𝐵) · 𝐶) = (𝐴 · (𝐵 · 𝐶))
 
Theoremadddii 7037 Distributive law (left-distributivity). (Contributed by NM, 23-Nov-1994.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ    &   𝐶 ∈ ℂ       (𝐴 · (𝐵 + 𝐶)) = ((𝐴 · 𝐵) + (𝐴 · 𝐶))
 
Theoremadddiri 7038 Distributive law (right-distributivity). (Contributed by NM, 16-Feb-1995.)
𝐴 ∈ ℂ    &   𝐵 ∈ ℂ    &   𝐶 ∈ ℂ       ((𝐴 + 𝐵) · 𝐶) = ((𝐴 · 𝐶) + (𝐵 · 𝐶))
 
Theoremrecni 7039 A real number is a complex number. (Contributed by NM, 1-Mar-1995.)
𝐴 ∈ ℝ       𝐴 ∈ ℂ
 
Theoremreaddcli 7040 Closure law for addition of reals. (Contributed by NM, 17-Jan-1997.)
𝐴 ∈ ℝ    &   𝐵 ∈ ℝ       (𝐴 + 𝐵) ∈ ℝ
 
Theoremremulcli 7041 Closure law for multiplication of reals. (Contributed by NM, 17-Jan-1997.)
𝐴 ∈ ℝ    &   𝐵 ∈ ℝ       (𝐴 · 𝐵) ∈ ℝ
 
Theorem1red 7042 1 is an real number, deductive form (common case). (Contributed by David A. Wheeler, 6-Dec-2018.)
(𝜑 → 1 ∈ ℝ)
 
Theorem1cnd 7043 1 is a complex number, deductive form (common case). (Contributed by David A. Wheeler, 6-Dec-2018.)
(𝜑 → 1 ∈ ℂ)
 
Theoremmulid1d 7044 Identity law for multiplication. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)       (𝜑 → (𝐴 · 1) = 𝐴)
 
Theoremmulid2d 7045 Identity law for multiplication. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)       (𝜑 → (1 · 𝐴) = 𝐴)
 
Theoremaddcld 7046 Closure law for addition. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)       (𝜑 → (𝐴 + 𝐵) ∈ ℂ)
 
Theoremmulcld 7047 Closure law for multiplication. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)       (𝜑 → (𝐴 · 𝐵) ∈ ℂ)
 
Theoremmulcomd 7048 Commutative law for multiplication. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)       (𝜑 → (𝐴 · 𝐵) = (𝐵 · 𝐴))
 
Theoremaddassd 7049 Associative law for addition. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐶 ∈ ℂ)       (𝜑 → ((𝐴 + 𝐵) + 𝐶) = (𝐴 + (𝐵 + 𝐶)))
 
Theoremmulassd 7050 Associative law for multiplication. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐶 ∈ ℂ)       (𝜑 → ((𝐴 · 𝐵) · 𝐶) = (𝐴 · (𝐵 · 𝐶)))
 
Theoremadddid 7051 Distributive law (left-distributivity). (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐶 ∈ ℂ)       (𝜑 → (𝐴 · (𝐵 + 𝐶)) = ((𝐴 · 𝐵) + (𝐴 · 𝐶)))
 
Theoremadddird 7052 Distributive law (right-distributivity). (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐶 ∈ ℂ)       (𝜑 → ((𝐴 + 𝐵) · 𝐶) = ((𝐴 · 𝐶) + (𝐵 · 𝐶)))
 
Theoremjoinlmuladdmuld 7053 Join AB+CB into (A+C) on LHS. (Contributed by David A. Wheeler, 26-Oct-2019.)
(𝜑𝐴 ∈ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐶 ∈ ℂ)    &   (𝜑 → ((𝐴 · 𝐵) + (𝐶 · 𝐵)) = 𝐷)       (𝜑 → ((𝐴 + 𝐶) · 𝐵) = 𝐷)
 
Theoremrecnd 7054 Deduction from real number to complex number. (Contributed by NM, 26-Oct-1999.)
(𝜑𝐴 ∈ ℝ)       (𝜑𝐴 ∈ ℂ)
 
Theoremreaddcld 7055 Closure law for addition of reals. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℝ)    &   (𝜑𝐵 ∈ ℝ)       (𝜑 → (𝐴 + 𝐵) ∈ ℝ)
 
Theoremremulcld 7056 Closure law for multiplication of reals. (Contributed by Mario Carneiro, 27-May-2016.)
(𝜑𝐴 ∈ ℝ)    &   (𝜑𝐵 ∈ ℝ)       (𝜑 → (𝐴 · 𝐵) ∈ ℝ)
 
3.2.2  Infinity and the extended real number system
 
Syntaxcpnf 7057 Plus infinity.
class +∞
 
Syntaxcmnf 7058 Minus infinity.
class -∞
 
Syntaxcxr 7059 The set of extended reals (includes plus and minus infinity).
class *
 
Syntaxclt 7060 'Less than' predicate (extended to include the extended reals).
class <
 
Syntaxcle 7061 Extend wff notation to include the 'less than or equal to' relation.
class
 
Definitiondf-pnf 7062 Define plus infinity. Note that the definition is arbitrary, requiring only that +∞ be a set not in and different from -∞ (df-mnf 7063). We use 𝒫 to make it independent of the construction of , and Cantor's Theorem will show that it is different from any member of and therefore . See pnfnre 7067 and mnfnre 7068, and we'll also be able to prove +∞ ≠ -∞.

A simpler possibility is to define +∞ as and -∞ as {ℂ}, but that approach requires the Axiom of Regularity to show that +∞ and -∞ are different from each other and from all members of . (Contributed by NM, 13-Oct-2005.) (New usage is discouraged.)

+∞ = 𝒫
 
Definitiondf-mnf 7063 Define minus infinity as the power set of plus infinity. Note that the definition is arbitrary, requiring only that -∞ be a set not in and different from +∞ (see mnfnre 7068). (Contributed by NM, 13-Oct-2005.) (New usage is discouraged.)
-∞ = 𝒫 +∞
 
Definitiondf-xr 7064 Define the set of extended reals that includes plus and minus infinity. Definition 12-3.1 of [Gleason] p. 173. (Contributed by NM, 13-Oct-2005.)
* = (ℝ ∪ {+∞, -∞})
 
Definitiondf-ltxr 7065* Define 'less than' on the set of extended reals. Definition 12-3.1 of [Gleason] p. 173. Note that in our postulates for complex numbers, < is primitive and not necessarily a relation on . (Contributed by NM, 13-Oct-2005.)
< = ({⟨𝑥, 𝑦⟩ ∣ (𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ ∧ 𝑥 < 𝑦)} ∪ (((ℝ ∪ {-∞}) × {+∞}) ∪ ({-∞} × ℝ)))
 
Definitiondf-le 7066 Define 'less than or equal to' on the extended real subset of complex numbers. (Contributed by NM, 13-Oct-2005.)
≤ = ((ℝ* × ℝ*) ∖ < )
 
Theorempnfnre 7067 Plus infinity is not a real number. (Contributed by NM, 13-Oct-2005.)
+∞ ∉ ℝ
 
Theoremmnfnre 7068 Minus infinity is not a real number. (Contributed by NM, 13-Oct-2005.)
-∞ ∉ ℝ
 
Theoremressxr 7069 The standard reals are a subset of the extended reals. (Contributed by NM, 14-Oct-2005.)
ℝ ⊆ ℝ*
 
Theoremrexpssxrxp 7070 The Cartesian product of standard reals are a subset of the Cartesian product of extended reals (common case). (Contributed by David A. Wheeler, 8-Dec-2018.)
(ℝ × ℝ) ⊆ (ℝ* × ℝ*)
 
Theoremrexr 7071 A standard real is an extended real. (Contributed by NM, 14-Oct-2005.)
(𝐴 ∈ ℝ → 𝐴 ∈ ℝ*)
 
Theorem0xr 7072 Zero is an extended real. (Contributed by Mario Carneiro, 15-Jun-2014.)
0 ∈ ℝ*
 
Theoremrenepnf 7073 No (finite) real equals plus infinity. (Contributed by NM, 14-Oct-2005.) (Proof shortened by Andrew Salmon, 19-Nov-2011.)
(𝐴 ∈ ℝ → 𝐴 ≠ +∞)
 
Theoremrenemnf 7074 No real equals minus infinity. (Contributed by NM, 14-Oct-2005.) (Proof shortened by Andrew Salmon, 19-Nov-2011.)
(𝐴 ∈ ℝ → 𝐴 ≠ -∞)
 
Theoremrexrd 7075 A standard real is an extended real. (Contributed by Mario Carneiro, 28-May-2016.)
(𝜑𝐴 ∈ ℝ)       (𝜑𝐴 ∈ ℝ*)
 
Theoremrenepnfd 7076 No (finite) real equals plus infinity. (Contributed by Mario Carneiro, 28-May-2016.)
(𝜑𝐴 ∈ ℝ)       (𝜑𝐴 ≠ +∞)
 
Theoremrenemnfd 7077 No real equals minus infinity. (Contributed by Mario Carneiro, 28-May-2016.)
(𝜑𝐴 ∈ ℝ)       (𝜑𝐴 ≠ -∞)
 
Theoremrexri 7078 A standard real is an extended real (inference form.) (Contributed by David Moews, 28-Feb-2017.)
𝐴 ∈ ℝ       𝐴 ∈ ℝ*
 
Theoremrenfdisj 7079 The reals and the infinities are disjoint. (Contributed by NM, 25-Oct-2005.) (Proof shortened by Andrew Salmon, 19-Nov-2011.)
(ℝ ∩ {+∞, -∞}) = ∅
 
Theoremltrelxr 7080 'Less than' is a relation on extended reals. (Contributed by Mario Carneiro, 28-Apr-2015.)
< ⊆ (ℝ* × ℝ*)
 
Theoremltrel 7081 'Less than' is a relation. (Contributed by NM, 14-Oct-2005.)
Rel <
 
Theoremlerelxr 7082 'Less than or equal' is a relation on extended reals. (Contributed by Mario Carneiro, 28-Apr-2015.)
≤ ⊆ (ℝ* × ℝ*)
 
Theoremlerel 7083 'Less or equal to' is a relation. (Contributed by FL, 2-Aug-2009.) (Revised by Mario Carneiro, 28-Apr-2015.)
Rel ≤
 
Theoremxrlenlt 7084 'Less than or equal to' expressed in terms of 'less than', for extended reals. (Contributed by NM, 14-Oct-2005.)
((𝐴 ∈ ℝ*𝐵 ∈ ℝ*) → (𝐴𝐵 ↔ ¬ 𝐵 < 𝐴))
 
Theoremltxrlt 7085 The standard less-than < and the extended real less-than < are identical when restricted to the non-extended reals . (Contributed by NM, 13-Oct-2005.) (Revised by Mario Carneiro, 28-Apr-2015.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 < 𝐵𝐴 < 𝐵))
 
3.2.3  Restate the ordering postulates with extended real "less than"
 
Theoremaxltirr 7086 Real number less-than is irreflexive. Axiom for real and complex numbers, derived from set theory. This restates ax-pre-ltirr 6996 with ordering on the extended reals. New proofs should use ltnr 7095 instead for naming consistency. (New usage is discouraged.) (Contributed by Jim Kingdon, 15-Jan-2020.)
(𝐴 ∈ ℝ → ¬ 𝐴 < 𝐴)
 
Theoremaxltwlin 7087 Real number less-than is weakly linear. Axiom for real and complex numbers, derived from set theory. This restates ax-pre-ltwlin 6997 with ordering on the extended reals. (Contributed by Jim Kingdon, 15-Jan-2020.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → (𝐴 < 𝐵 → (𝐴 < 𝐶𝐶 < 𝐵)))
 
Theoremaxlttrn 7088 Ordering on reals is transitive. Axiom for real and complex numbers, derived from set theory. This restates ax-pre-lttrn 6998 with ordering on the extended reals. New proofs should use lttr 7092 instead for naming consistency. (New usage is discouraged.) (Contributed by NM, 13-Oct-2005.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → ((𝐴 < 𝐵𝐵 < 𝐶) → 𝐴 < 𝐶))
 
Theoremaxltadd 7089 Ordering property of addition on reals. Axiom for real and complex numbers, derived from set theory. (This restates ax-pre-ltadd 7000 with ordering on the extended reals.) (Contributed by NM, 13-Oct-2005.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → (𝐴 < 𝐵 → (𝐶 + 𝐴) < (𝐶 + 𝐵)))
 
Theoremaxapti 7090 Apartness of reals is tight. Axiom for real and complex numbers, derived from set theory. (This restates ax-pre-apti 6999 with ordering on the extended reals.) (Contributed by Jim Kingdon, 29-Jan-2020.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ ¬ (𝐴 < 𝐵𝐵 < 𝐴)) → 𝐴 = 𝐵)
 
Theoremaxmulgt0 7091 The product of two positive reals is positive. Axiom for real and complex numbers, derived from set theory. (This restates ax-pre-mulgt0 7001 with ordering on the extended reals.) (Contributed by NM, 13-Oct-2005.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → ((0 < 𝐴 ∧ 0 < 𝐵) → 0 < (𝐴 · 𝐵)))
 
3.2.4  Ordering on reals
 
Theoremlttr 7092 Alias for axlttrn 7088, for naming consistency with lttri 7122. New proofs should generally use this instead of ax-pre-lttrn 6998. (Contributed by NM, 10-Mar-2008.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → ((𝐴 < 𝐵𝐵 < 𝐶) → 𝐴 < 𝐶))
 
Theoremmulgt0 7093 The product of two positive numbers is positive. (Contributed by NM, 10-Mar-2008.)
(((𝐴 ∈ ℝ ∧ 0 < 𝐴) ∧ (𝐵 ∈ ℝ ∧ 0 < 𝐵)) → 0 < (𝐴 · 𝐵))
 
Theoremlenlt 7094 'Less than or equal to' expressed in terms of 'less than'. Part of definition 11.2.7(vi) of [HoTT], p. (varies). (Contributed by NM, 13-May-1999.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴𝐵 ↔ ¬ 𝐵 < 𝐴))
 
Theoremltnr 7095 'Less than' is irreflexive. (Contributed by NM, 18-Aug-1999.)
(𝐴 ∈ ℝ → ¬ 𝐴 < 𝐴)
 
Theoremltso 7096 'Less than' is a strict ordering. (Contributed by NM, 19-Jan-1997.)
< Or ℝ
 
Theoremgtso 7097 'Greater than' is a strict ordering. (Contributed by JJ, 11-Oct-2018.)
< Or ℝ
 
Theoremlttri3 7098 Tightness of real apartness. (Contributed by NM, 5-May-1999.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 = 𝐵 ↔ (¬ 𝐴 < 𝐵 ∧ ¬ 𝐵 < 𝐴)))
 
Theoremletri3 7099 Tightness of real apartness. (Contributed by NM, 14-May-1999.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 = 𝐵 ↔ (𝐴𝐵𝐵𝐴)))
 
Theoremltleletr 7100 Transitive law, weaker form of (𝐴 < 𝐵𝐵𝐶) → 𝐴 < 𝐶. (Contributed by AV, 14-Oct-2018.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → ((𝐴 < 𝐵𝐵𝐶) → 𝐴𝐶))
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10124
  Copyright terms: Public domain < Previous  Next >