| Intuitionistic Logic Explorer Theorem List (p. 153 of 159) | < Previous Next > | |
| Bad symbols? Try the
GIF version. |
||
|
Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
||
| Type | Label | Description |
|---|---|---|
| Statement | ||
| Theorem | logdivlti 15201 | The log𝑥 / 𝑥 function is strictly decreasing on the reals greater than e. (Contributed by Mario Carneiro, 14-Mar-2014.) |
| ⊢ (((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ ∧ e ≤ 𝐴) ∧ 𝐴 < 𝐵) → ((log‘𝐵) / 𝐵) < ((log‘𝐴) / 𝐴)) | ||
| Theorem | relogcld 15202 | Closure of the natural logarithm function. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) ⇒ ⊢ (𝜑 → (log‘𝐴) ∈ ℝ) | ||
| Theorem | reeflogd 15203 | Relationship between the natural logarithm function and the exponential function. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) ⇒ ⊢ (𝜑 → (exp‘(log‘𝐴)) = 𝐴) | ||
| Theorem | relogmuld 15204 | The natural logarithm of the product of two positive real numbers is the sum of natural logarithms. Property 2 of [Cohen] p. 301, restricted to natural logarithms. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ+) ⇒ ⊢ (𝜑 → (log‘(𝐴 · 𝐵)) = ((log‘𝐴) + (log‘𝐵))) | ||
| Theorem | relogdivd 15205 | The natural logarithm of the quotient of two positive real numbers is the difference of natural logarithms. Exercise 72(a) and Property 3 of [Cohen] p. 301, restricted to natural logarithms. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ+) ⇒ ⊢ (𝜑 → (log‘(𝐴 / 𝐵)) = ((log‘𝐴) − (log‘𝐵))) | ||
| Theorem | logled 15206 | Natural logarithm preserves ≤. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ+) ⇒ ⊢ (𝜑 → (𝐴 ≤ 𝐵 ↔ (log‘𝐴) ≤ (log‘𝐵))) | ||
| Theorem | relogefd 15207 | Relationship between the natural logarithm function and the exponential function. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ) ⇒ ⊢ (𝜑 → (log‘(exp‘𝐴)) = 𝐴) | ||
| Theorem | rplogcld 15208 | Closure of the logarithm function in the positive reals. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 1 < 𝐴) ⇒ ⊢ (𝜑 → (log‘𝐴) ∈ ℝ+) | ||
| Theorem | logge0d 15209 | The logarithm of a number greater than 1 is nonnegative. (Contributed by Mario Carneiro, 29-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 1 ≤ 𝐴) ⇒ ⊢ (𝜑 → 0 ≤ (log‘𝐴)) | ||
| Theorem | logge0b 15210 | The logarithm of a number is nonnegative iff the number is greater than or equal to 1. (Contributed by AV, 30-May-2020.) |
| ⊢ (𝐴 ∈ ℝ+ → (0 ≤ (log‘𝐴) ↔ 1 ≤ 𝐴)) | ||
| Theorem | loggt0b 15211 | The logarithm of a number is positive iff the number is greater than 1. (Contributed by AV, 30-May-2020.) |
| ⊢ (𝐴 ∈ ℝ+ → (0 < (log‘𝐴) ↔ 1 < 𝐴)) | ||
| Theorem | logle1b 15212 | The logarithm of a number is less than or equal to 1 iff the number is less than or equal to Euler's constant. (Contributed by AV, 30-May-2020.) |
| ⊢ (𝐴 ∈ ℝ+ → ((log‘𝐴) ≤ 1 ↔ 𝐴 ≤ e)) | ||
| Theorem | loglt1b 15213 | The logarithm of a number is less than 1 iff the number is less than Euler's constant. (Contributed by AV, 30-May-2020.) |
| ⊢ (𝐴 ∈ ℝ+ → ((log‘𝐴) < 1 ↔ 𝐴 < e)) | ||
| Theorem | rpcxpef 15214 | Value of the complex power function. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴)))) | ||
| Theorem | cxpexprp 15215 | Relate the complex power function to the integer power function. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℤ) → (𝐴↑𝑐𝐵) = (𝐴↑𝐵)) | ||
| Theorem | cxpexpnn 15216 | Relate the complex power function to the integer power function. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
| ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴↑𝑐𝐵) = (𝐴↑𝐵)) | ||
| Theorem | logcxp 15217 | Logarithm of a complex power. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ) → (log‘(𝐴↑𝑐𝐵)) = (𝐵 · (log‘𝐴))) | ||
| Theorem | rpcxp0 15218 | Value of the complex power function when the second argument is zero. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
| ⊢ (𝐴 ∈ ℝ+ → (𝐴↑𝑐0) = 1) | ||
| Theorem | rpcxp1 15219 | Value of the complex power function at one. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ (𝐴 ∈ ℝ+ → (𝐴↑𝑐1) = 𝐴) | ||
| Theorem | 1cxp 15220 | Value of the complex power function at one. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ (𝐴 ∈ ℂ → (1↑𝑐𝐴) = 1) | ||
| Theorem | ecxp 15221 | Write the exponential function as an exponent to the power e. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ (𝐴 ∈ ℂ → (e↑𝑐𝐴) = (exp‘𝐴)) | ||
| Theorem | rpcncxpcl 15222 | Closure of the complex power function. (Contributed by Jim Kingdon, 12-Jun-2024.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐𝐵) ∈ ℂ) | ||
| Theorem | rpcxpcl 15223 | Positive real closure of the complex power function. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ) → (𝐴↑𝑐𝐵) ∈ ℝ+) | ||
| Theorem | cxpap0 15224 | Complex exponentiation is apart from zero. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐𝐵) # 0) | ||
| Theorem | rpcxpadd 15225 | Sum of exponents law for complex exponentiation. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 13-Jun-2024.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → (𝐴↑𝑐(𝐵 + 𝐶)) = ((𝐴↑𝑐𝐵) · (𝐴↑𝑐𝐶))) | ||
| Theorem | rpcxpp1 15226 | Value of a nonzero complex number raised to a complex power plus one. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐(𝐵 + 1)) = ((𝐴↑𝑐𝐵) · 𝐴)) | ||
| Theorem | rpcxpneg 15227 | Value of a complex number raised to a negative power. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐-𝐵) = (1 / (𝐴↑𝑐𝐵))) | ||
| Theorem | rpcxpsub 15228 | Exponent subtraction law for complex exponentiation. (Contributed by Mario Carneiro, 22-Sep-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → (𝐴↑𝑐(𝐵 − 𝐶)) = ((𝐴↑𝑐𝐵) / (𝐴↑𝑐𝐶))) | ||
| Theorem | rpmulcxp 15229 | Complex exponentiation of a product. Proposition 10-4.2(c) of [Gleason] p. 135. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ+ ∧ 𝐶 ∈ ℂ) → ((𝐴 · 𝐵)↑𝑐𝐶) = ((𝐴↑𝑐𝐶) · (𝐵↑𝑐𝐶))) | ||
| Theorem | cxprec 15230 | Complex exponentiation of a reciprocal. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → ((1 / 𝐴)↑𝑐𝐵) = (1 / (𝐴↑𝑐𝐵))) | ||
| Theorem | rpdivcxp 15231 | Complex exponentiation of a quotient. (Contributed by Mario Carneiro, 8-Sep-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ+ ∧ 𝐶 ∈ ℂ) → ((𝐴 / 𝐵)↑𝑐𝐶) = ((𝐴↑𝑐𝐶) / (𝐵↑𝑐𝐶))) | ||
| Theorem | cxpmul 15232 | Product of exponents law for complex exponentiation. Proposition 10-4.2(b) of [Gleason] p. 135. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℂ) → (𝐴↑𝑐(𝐵 · 𝐶)) = ((𝐴↑𝑐𝐵)↑𝑐𝐶)) | ||
| Theorem | rpcxpmul2 15233 | Product of exponents law for complex exponentiation. Variation on cxpmul 15232 with more general conditions on 𝐴 and 𝐵 when 𝐶 is a nonnegative integer. (Contributed by Mario Carneiro, 9-Aug-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℕ0) → (𝐴↑𝑐(𝐵 · 𝐶)) = ((𝐴↑𝑐𝐵)↑𝐶)) | ||
| Theorem | rpcxproot 15234 | The complex power function allows us to write n-th roots via the idiom 𝐴↑𝑐(1 / 𝑁). (Contributed by Mario Carneiro, 6-May-2015.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝑁 ∈ ℕ) → ((𝐴↑𝑐(1 / 𝑁))↑𝑁) = 𝐴) | ||
| Theorem | abscxp 15235 | Absolute value of a power, when the base is real. (Contributed by Mario Carneiro, 15-Sep-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (abs‘(𝐴↑𝑐𝐵)) = (𝐴↑𝑐(ℜ‘𝐵))) | ||
| Theorem | cxplt 15236 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ (((𝐴 ∈ ℝ ∧ 1 < 𝐴) ∧ (𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ)) → (𝐵 < 𝐶 ↔ (𝐴↑𝑐𝐵) < (𝐴↑𝑐𝐶))) | ||
| Theorem | cxple 15237 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 2-Aug-2014.) |
| ⊢ (((𝐴 ∈ ℝ ∧ 1 < 𝐴) ∧ (𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ)) → (𝐵 ≤ 𝐶 ↔ (𝐴↑𝑐𝐵) ≤ (𝐴↑𝑐𝐶))) | ||
| Theorem | rpcxple2 15238 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 8-Sep-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ+ ∧ 𝐶 ∈ ℝ+) → (𝐴 ≤ 𝐵 ↔ (𝐴↑𝑐𝐶) ≤ (𝐵↑𝑐𝐶))) | ||
| Theorem | rpcxplt2 15239 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 15-Sep-2014.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ+ ∧ 𝐶 ∈ ℝ+) → (𝐴 < 𝐵 ↔ (𝐴↑𝑐𝐶) < (𝐵↑𝑐𝐶))) | ||
| Theorem | cxplt3 15240 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 2-May-2016.) |
| ⊢ (((𝐴 ∈ ℝ+ ∧ 𝐴 < 1) ∧ (𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ)) → (𝐵 < 𝐶 ↔ (𝐴↑𝑐𝐶) < (𝐴↑𝑐𝐵))) | ||
| Theorem | cxple3 15241 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 2-May-2016.) |
| ⊢ (((𝐴 ∈ ℝ+ ∧ 𝐴 < 1) ∧ (𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ)) → (𝐵 ≤ 𝐶 ↔ (𝐴↑𝑐𝐶) ≤ (𝐴↑𝑐𝐵))) | ||
| Theorem | rpcxpsqrt 15242 | The exponential function with exponent 1 / 2 exactly matches the square root function, and thus serves as a suitable generalization to other 𝑛-th roots and irrational roots. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 16-Jun-2024.) |
| ⊢ (𝐴 ∈ ℝ+ → (𝐴↑𝑐(1 / 2)) = (√‘𝐴)) | ||
| Theorem | logsqrt 15243 | Logarithm of a square root. (Contributed by Mario Carneiro, 5-May-2016.) |
| ⊢ (𝐴 ∈ ℝ+ → (log‘(√‘𝐴)) = ((log‘𝐴) / 2)) | ||
| Theorem | rpcxp0d 15244 | Value of the complex power function when the second argument is zero. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) ⇒ ⊢ (𝜑 → (𝐴↑𝑐0) = 1) | ||
| Theorem | rpcxp1d 15245 | Value of the complex power function at one. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) ⇒ ⊢ (𝜑 → (𝐴↑𝑐1) = 𝐴) | ||
| Theorem | 1cxpd 15246 | Value of the complex power function at one. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℂ) ⇒ ⊢ (𝜑 → (1↑𝑐𝐴) = 1) | ||
| Theorem | rpcncxpcld 15247 | Closure of the complex power function. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℂ) ⇒ ⊢ (𝜑 → (𝐴↑𝑐𝐵) ∈ ℂ) | ||
| Theorem | cxpltd 15248 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 1 < 𝐴) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℝ) ⇒ ⊢ (𝜑 → (𝐵 < 𝐶 ↔ (𝐴↑𝑐𝐵) < (𝐴↑𝑐𝐶))) | ||
| Theorem | cxpled 15249 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 1 < 𝐴) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℝ) ⇒ ⊢ (𝜑 → (𝐵 ≤ 𝐶 ↔ (𝐴↑𝑐𝐵) ≤ (𝐴↑𝑐𝐶))) | ||
| Theorem | rpcxpsqrtth 15250 | Square root theorem over the complex numbers for the complex power function. Compare with resqrtth 11213. (Contributed by AV, 23-Dec-2022.) |
| ⊢ (𝐴 ∈ ℝ+ → ((√‘𝐴)↑𝑐2) = 𝐴) | ||
| Theorem | cxprecd 15251 | Complex exponentiation of a reciprocal. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℂ) ⇒ ⊢ (𝜑 → ((1 / 𝐴)↑𝑐𝐵) = (1 / (𝐴↑𝑐𝐵))) | ||
| Theorem | rpcxpmul2d 15252 | Product of exponents law for complex exponentiation. Variation on cxpmul 15232 with more general conditions on 𝐴 and 𝐵 when 𝐶 is a nonnegative integer. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 𝐶 ∈ ℕ0) ⇒ ⊢ (𝜑 → (𝐴↑𝑐(𝐵 · 𝐶)) = ((𝐴↑𝑐𝐵)↑𝐶)) | ||
| Theorem | rpcxpcld 15253 | Positive real closure of the complex power function. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ) ⇒ ⊢ (𝜑 → (𝐴↑𝑐𝐵) ∈ ℝ+) | ||
| Theorem | logcxpd 15254 | Logarithm of a complex power. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ) ⇒ ⊢ (𝜑 → (log‘(𝐴↑𝑐𝐵)) = (𝐵 · (log‘𝐴))) | ||
| Theorem | cxplt3d 15255 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 1) & ⊢ (𝜑 → 𝐶 ∈ ℝ) ⇒ ⊢ (𝜑 → (𝐵 < 𝐶 ↔ (𝐴↑𝑐𝐶) < (𝐴↑𝑐𝐵))) | ||
| Theorem | cxple3d 15256 | Ordering property for complex exponentiation. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 1) & ⊢ (𝜑 → 𝐶 ∈ ℝ) ⇒ ⊢ (𝜑 → (𝐵 ≤ 𝐶 ↔ (𝐴↑𝑐𝐶) ≤ (𝐴↑𝑐𝐵))) | ||
| Theorem | cxpmuld 15257 | Product of exponents law for complex exponentiation. Proposition 10-4.2(b) of [Gleason] p. 135. (Contributed by Mario Carneiro, 30-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℂ) ⇒ ⊢ (𝜑 → (𝐴↑𝑐(𝐵 · 𝐶)) = ((𝐴↑𝑐𝐵)↑𝑐𝐶)) | ||
| Theorem | cxpcom 15258 | Commutative law for real exponentiation. (Contributed by AV, 29-Dec-2022.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ) → ((𝐴↑𝑐𝐵)↑𝑐𝐶) = ((𝐴↑𝑐𝐶)↑𝑐𝐵)) | ||
| Theorem | apcxp2 15259 | Apartness and real exponentiation. (Contributed by Jim Kingdon, 10-Jul-2024.) |
| ⊢ (((𝐴 ∈ ℝ+ ∧ 𝐴 # 1) ∧ (𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ)) → (𝐵 # 𝐶 ↔ (𝐴↑𝑐𝐵) # (𝐴↑𝑐𝐶))) | ||
| Theorem | rpabscxpbnd 15260 | Bound on the absolute value of a complex power. (Contributed by Mario Carneiro, 15-Sep-2014.) (Revised by Jim Kingdon, 19-Jun-2024.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 0 < (ℜ‘𝐵)) & ⊢ (𝜑 → 𝑀 ∈ ℝ) & ⊢ (𝜑 → (abs‘𝐴) ≤ 𝑀) ⇒ ⊢ (𝜑 → (abs‘(𝐴↑𝑐𝐵)) ≤ ((𝑀↑𝑐(ℜ‘𝐵)) · (exp‘((abs‘𝐵) · π)))) | ||
| Theorem | ltexp2 15261 | Ordering law for exponentiation. (Contributed by NM, 2-Aug-2006.) (Revised by Mario Carneiro, 5-Jun-2014.) |
| ⊢ (((𝐴 ∈ ℝ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ 1 < 𝐴) → (𝑀 < 𝑁 ↔ (𝐴↑𝑀) < (𝐴↑𝑁))) | ||
| Theorem | ltexp2d 15262 | Ordering relationship for exponentiation. (Contributed by Mario Carneiro, 28-May-2016.) |
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℤ) & ⊢ (𝜑 → 1 < 𝐴) ⇒ ⊢ (𝜑 → (𝑀 < 𝑁 ↔ (𝐴↑𝑀) < (𝐴↑𝑁))) | ||
Define "log using an arbitrary base" function and then prove some of its properties. As with df-relog 15178 this is for real logarithms rather than complex logarithms. Metamath doesn't care what letters are used to represent classes. Usually classes begin with the letter "A", but here we use "B" and "X" to more clearly distinguish between "base" and "other parameter of log". There are different ways this could be defined in Metamath. The approach used here is intentionally similar to existing 2-parameter Metamath functions (operations): (𝐵 logb 𝑋) where 𝐵 is the base and 𝑋 is the argument of the logarithm function. An alternative would be to support the notational form (( logb ‘𝐵)‘𝑋); that looks a little more like traditional notation. | ||
| Syntax | clogb 15263 | Extend class notation to include the logarithm generalized to an arbitrary base. |
| class logb | ||
| Definition | df-logb 15264* | Define the logb operator. This is the logarithm generalized to an arbitrary base. It can be used as (𝐵 logb 𝑋) for "log base B of X". In the most common traditional notation, base B is a subscript of "log". The definition will only be useful where 𝑥 is a positive real apart from one and where 𝑦 is a positive real, so the choice of (ℂ ∖ {0, 1}) and (ℂ ∖ {0}) is somewhat arbitrary (we adopt the definition used in set.mm). (Contributed by David A. Wheeler, 21-Jan-2017.) |
| ⊢ logb = (𝑥 ∈ (ℂ ∖ {0, 1}), 𝑦 ∈ (ℂ ∖ {0}) ↦ ((log‘𝑦) / (log‘𝑥))) | ||
| Theorem | rplogbval 15265 | Define the value of the logb function, the logarithm generalized to an arbitrary base, when used as infix. Most Metamath statements select variables in order of their use, but to make the order clearer we use "B" for base and "X" for the argument of the logarithm function here. (Contributed by David A. Wheeler, 21-Jan-2017.) (Revised by Jim Kingdon, 3-Jul-2024.) |
| ⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1 ∧ 𝑋 ∈ ℝ+) → (𝐵 logb 𝑋) = ((log‘𝑋) / (log‘𝐵))) | ||
| Theorem | rplogbcl 15266 | General logarithm closure. (Contributed by David A. Wheeler, 17-Jul-2017.) |
| ⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1 ∧ 𝑋 ∈ ℝ+) → (𝐵 logb 𝑋) ∈ ℝ) | ||
| Theorem | rplogbid1 15267 | General logarithm is 1 when base and arg match. Property 1(a) of [Cohen4] p. 361. (Contributed by Stefan O'Rear, 19-Sep-2014.) (Revised by David A. Wheeler, 22-Jul-2017.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ 𝐴 # 1) → (𝐴 logb 𝐴) = 1) | ||
| Theorem | rplogb1 15268 | The logarithm of 1 to an arbitrary base 𝐵 is 0. Property 1(b) of [Cohen4] p. 361. See log1 15186. (Contributed by Stefan O'Rear, 19-Sep-2014.) (Revised by Thierry Arnoux, 27-Sep-2017.) |
| ⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) → (𝐵 logb 1) = 0) | ||
| Theorem | rpelogb 15269 | The general logarithm of a number to the base being Euler's constant is the natural logarithm of the number. Put another way, using e as the base in logb is the same as log. Definition in [Cohen4] p. 352. (Contributed by David A. Wheeler, 17-Oct-2017.) (Revised by David A. Wheeler and AV, 16-Jun-2020.) |
| ⊢ (𝐴 ∈ ℝ+ → (e logb 𝐴) = (log‘𝐴)) | ||
| Theorem | rplogbchbase 15270 | Change of base for logarithms. Property in [Cohen4] p. 367. (Contributed by AV, 11-Jun-2020.) |
| ⊢ (((𝐴 ∈ ℝ+ ∧ 𝐴 # 1) ∧ (𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ 𝑋 ∈ ℝ+) → (𝐴 logb 𝑋) = ((𝐵 logb 𝑋) / (𝐵 logb 𝐴))) | ||
| Theorem | relogbval 15271 | Value of the general logarithm with integer base. (Contributed by Thierry Arnoux, 27-Sep-2017.) |
| ⊢ ((𝐵 ∈ (ℤ≥‘2) ∧ 𝑋 ∈ ℝ+) → (𝐵 logb 𝑋) = ((log‘𝑋) / (log‘𝐵))) | ||
| Theorem | relogbzcl 15272 | Closure of the general logarithm with integer base on positive reals. (Contributed by Thierry Arnoux, 27-Sep-2017.) (Proof shortened by AV, 9-Jun-2020.) |
| ⊢ ((𝐵 ∈ (ℤ≥‘2) ∧ 𝑋 ∈ ℝ+) → (𝐵 logb 𝑋) ∈ ℝ) | ||
| Theorem | rplogbreexp 15273 | Power law for the general logarithm for real powers: The logarithm of a positive real number to the power of a real number is equal to the product of the exponent and the logarithm of the base of the power. Property 4 of [Cohen4] p. 361. (Contributed by AV, 9-Jun-2020.) |
| ⊢ (((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ 𝐶 ∈ ℝ+ ∧ 𝐸 ∈ ℝ) → (𝐵 logb (𝐶↑𝑐𝐸)) = (𝐸 · (𝐵 logb 𝐶))) | ||
| Theorem | rplogbzexp 15274 | Power law for the general logarithm for integer powers: The logarithm of a positive real number to the power of an integer is equal to the product of the exponent and the logarithm of the base of the power. (Contributed by Stefan O'Rear, 19-Sep-2014.) (Revised by AV, 9-Jun-2020.) |
| ⊢ (((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ 𝐶 ∈ ℝ+ ∧ 𝑁 ∈ ℤ) → (𝐵 logb (𝐶↑𝑁)) = (𝑁 · (𝐵 logb 𝐶))) | ||
| Theorem | rprelogbmul 15275 | The logarithm of the product of two positive real numbers is the sum of logarithms. Property 2 of [Cohen4] p. 361. (Contributed by Stefan O'Rear, 19-Sep-2014.) (Revised by AV, 29-May-2020.) |
| ⊢ (((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ (𝐴 ∈ ℝ+ ∧ 𝐶 ∈ ℝ+)) → (𝐵 logb (𝐴 · 𝐶)) = ((𝐵 logb 𝐴) + (𝐵 logb 𝐶))) | ||
| Theorem | rprelogbmulexp 15276 | The logarithm of the product of a positive real and a positive real number to the power of a real number is the sum of the logarithm of the first real number and the scaled logarithm of the second real number. (Contributed by AV, 29-May-2020.) |
| ⊢ (((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ (𝐴 ∈ ℝ+ ∧ 𝐶 ∈ ℝ+ ∧ 𝐸 ∈ ℝ)) → (𝐵 logb (𝐴 · (𝐶↑𝑐𝐸))) = ((𝐵 logb 𝐴) + (𝐸 · (𝐵 logb 𝐶)))) | ||
| Theorem | rprelogbdiv 15277 | The logarithm of the quotient of two positive real numbers is the difference of logarithms. Property 3 of [Cohen4] p. 361. (Contributed by AV, 29-May-2020.) |
| ⊢ (((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ (𝐴 ∈ ℝ+ ∧ 𝐶 ∈ ℝ+)) → (𝐵 logb (𝐴 / 𝐶)) = ((𝐵 logb 𝐴) − (𝐵 logb 𝐶))) | ||
| Theorem | relogbexpap 15278 | Identity law for general logarithm: the logarithm of a power to the base is the exponent. Property 6 of [Cohen4] p. 361. (Contributed by Stefan O'Rear, 19-Sep-2014.) (Revised by AV, 9-Jun-2020.) |
| ⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1 ∧ 𝑀 ∈ ℤ) → (𝐵 logb (𝐵↑𝑀)) = 𝑀) | ||
| Theorem | nnlogbexp 15279 | Identity law for general logarithm with integer base. (Contributed by Stefan O'Rear, 19-Sep-2014.) (Revised by Thierry Arnoux, 27-Sep-2017.) |
| ⊢ ((𝐵 ∈ (ℤ≥‘2) ∧ 𝑀 ∈ ℤ) → (𝐵 logb (𝐵↑𝑀)) = 𝑀) | ||
| Theorem | logbrec 15280 | Logarithm of a reciprocal changes sign. Particular case of Property 3 of [Cohen4] p. 361. (Contributed by Thierry Arnoux, 27-Sep-2017.) |
| ⊢ ((𝐵 ∈ (ℤ≥‘2) ∧ 𝐴 ∈ ℝ+) → (𝐵 logb (1 / 𝐴)) = -(𝐵 logb 𝐴)) | ||
| Theorem | logbleb 15281 | The general logarithm function is monotone/increasing. See logleb 15195. (Contributed by Stefan O'Rear, 19-Oct-2014.) (Revised by AV, 31-May-2020.) |
| ⊢ ((𝐵 ∈ (ℤ≥‘2) ∧ 𝑋 ∈ ℝ+ ∧ 𝑌 ∈ ℝ+) → (𝑋 ≤ 𝑌 ↔ (𝐵 logb 𝑋) ≤ (𝐵 logb 𝑌))) | ||
| Theorem | logblt 15282 | The general logarithm function is strictly monotone/increasing. Property 2 of [Cohen4] p. 377. See logltb 15194. (Contributed by Stefan O'Rear, 19-Oct-2014.) (Revised by Thierry Arnoux, 27-Sep-2017.) |
| ⊢ ((𝐵 ∈ (ℤ≥‘2) ∧ 𝑋 ∈ ℝ+ ∧ 𝑌 ∈ ℝ+) → (𝑋 < 𝑌 ↔ (𝐵 logb 𝑋) < (𝐵 logb 𝑌))) | ||
| Theorem | rplogbcxp 15283 | Identity law for the general logarithm for real numbers. (Contributed by AV, 22-May-2020.) |
| ⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1 ∧ 𝑋 ∈ ℝ) → (𝐵 logb (𝐵↑𝑐𝑋)) = 𝑋) | ||
| Theorem | rpcxplogb 15284 | Identity law for the general logarithm. (Contributed by AV, 22-May-2020.) |
| ⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1 ∧ 𝑋 ∈ ℝ+) → (𝐵↑𝑐(𝐵 logb 𝑋)) = 𝑋) | ||
| Theorem | relogbcxpbap 15285 | The logarithm is the inverse of the exponentiation. Observation in [Cohen4] p. 348. (Contributed by AV, 11-Jun-2020.) |
| ⊢ (((𝐵 ∈ ℝ+ ∧ 𝐵 # 1) ∧ 𝑋 ∈ ℝ+ ∧ 𝑌 ∈ ℝ) → ((𝐵 logb 𝑋) = 𝑌 ↔ (𝐵↑𝑐𝑌) = 𝑋)) | ||
| Theorem | logbgt0b 15286 | The logarithm of a positive real number to a real base greater than 1 is positive iff the number is greater than 1. (Contributed by AV, 29-Dec-2022.) |
| ⊢ ((𝐴 ∈ ℝ+ ∧ (𝐵 ∈ ℝ+ ∧ 1 < 𝐵)) → (0 < (𝐵 logb 𝐴) ↔ 1 < 𝐴)) | ||
| Theorem | logbgcd1irr 15287 | The logarithm of an integer greater than 1 to an integer base greater than 1 is not rational if the argument and the base are relatively prime. For example, (2 logb 9) ∈ (ℝ ∖ ℚ). (Contributed by AV, 29-Dec-2022.) |
| ⊢ ((𝑋 ∈ (ℤ≥‘2) ∧ 𝐵 ∈ (ℤ≥‘2) ∧ (𝑋 gcd 𝐵) = 1) → (𝐵 logb 𝑋) ∈ (ℝ ∖ ℚ)) | ||
| Theorem | logbgcd1irraplemexp 15288 | Lemma for logbgcd1irrap 15290. Apartness of 𝑋↑𝑁 and 𝐵↑𝑀. (Contributed by Jim Kingdon, 11-Jul-2024.) |
| ⊢ (𝜑 → 𝑋 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐵 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → (𝑋 gcd 𝐵) = 1) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℕ) ⇒ ⊢ (𝜑 → (𝑋↑𝑁) # (𝐵↑𝑀)) | ||
| Theorem | logbgcd1irraplemap 15289 | Lemma for logbgcd1irrap 15290. The result, with the rational number expressed as numerator and denominator. (Contributed by Jim Kingdon, 9-Jul-2024.) |
| ⊢ (𝜑 → 𝑋 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐵 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → (𝑋 gcd 𝐵) = 1) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℕ) ⇒ ⊢ (𝜑 → (𝐵 logb 𝑋) # (𝑀 / 𝑁)) | ||
| Theorem | logbgcd1irrap 15290 | The logarithm of an integer greater than 1 to an integer base greater than 1 is irrational (in the sense of being apart from any rational number) if the argument and the base are relatively prime. For example, (2 logb 9) # 𝑄 where 𝑄 is rational. (Contributed by AV, 29-Dec-2022.) |
| ⊢ (((𝑋 ∈ (ℤ≥‘2) ∧ 𝐵 ∈ (ℤ≥‘2)) ∧ ((𝑋 gcd 𝐵) = 1 ∧ 𝑄 ∈ ℚ)) → (𝐵 logb 𝑋) # 𝑄) | ||
| Theorem | 2logb9irr 15291 | Example for logbgcd1irr 15287. The logarithm of nine to base two is not rational. Also see 2logb9irrap 15297 which says that it is irrational (in the sense of being apart from any rational number). (Contributed by AV, 29-Dec-2022.) |
| ⊢ (2 logb 9) ∈ (ℝ ∖ ℚ) | ||
| Theorem | logbprmirr 15292 | The logarithm of a prime to a different prime base is not rational. For example, (2 logb 3) ∈ (ℝ ∖ ℚ) (see 2logb3irr 15293). (Contributed by AV, 31-Dec-2022.) |
| ⊢ ((𝑋 ∈ ℙ ∧ 𝐵 ∈ ℙ ∧ 𝑋 ≠ 𝐵) → (𝐵 logb 𝑋) ∈ (ℝ ∖ ℚ)) | ||
| Theorem | 2logb3irr 15293 | Example for logbprmirr 15292. The logarithm of three to base two is not rational. (Contributed by AV, 31-Dec-2022.) |
| ⊢ (2 logb 3) ∈ (ℝ ∖ ℚ) | ||
| Theorem | 2logb9irrALT 15294 | Alternate proof of 2logb9irr 15291: The logarithm of nine to base two is not rational. (Contributed by AV, 31-Dec-2022.) (Proof modification is discouraged.) (New usage is discouraged.) |
| ⊢ (2 logb 9) ∈ (ℝ ∖ ℚ) | ||
| Theorem | sqrt2cxp2logb9e3 15295 | The square root of two to the power of the logarithm of nine to base two is three. (√‘2) and (2 logb 9) are not rational (see sqrt2irr0 12357 resp. 2logb9irr 15291), satisfying the statement in 2irrexpq 15296. (Contributed by AV, 29-Dec-2022.) |
| ⊢ ((√‘2)↑𝑐(2 logb 9)) = 3 | ||
| Theorem | 2irrexpq 15296* |
There exist real numbers 𝑎 and 𝑏 which are not rational
such
that (𝑎↑𝑏) is rational. Statement in the
Metamath book, section
1.1.5, footnote 27 on page 17, and the "constructive proof"
for theorem
1.2 of [Bauer], p. 483. This is a
constructive proof because it is
based on two explicitly named non-rational numbers (√‘2) and
(2 logb 9), see sqrt2irr0 12357, 2logb9irr 15291 and
sqrt2cxp2logb9e3 15295. Therefore, this proof is acceptable/usable
in
intuitionistic logic.
For a theorem which is the same but proves that 𝑎 and 𝑏 are irrational (in the sense of being apart from any rational number), see 2irrexpqap 15298. (Contributed by AV, 23-Dec-2022.) |
| ⊢ ∃𝑎 ∈ (ℝ ∖ ℚ)∃𝑏 ∈ (ℝ ∖ ℚ)(𝑎↑𝑐𝑏) ∈ ℚ | ||
| Theorem | 2logb9irrap 15297 | Example for logbgcd1irrap 15290. The logarithm of nine to base two is irrational (in the sense of being apart from any rational number). (Contributed by Jim Kingdon, 12-Jul-2024.) |
| ⊢ (𝑄 ∈ ℚ → (2 logb 9) # 𝑄) | ||
| Theorem | 2irrexpqap 15298* | There exist real numbers 𝑎 and 𝑏 which are irrational (in the sense of being apart from any rational number) such that (𝑎↑𝑏) is rational. Statement in the Metamath book, section 1.1.5, footnote 27 on page 17, and the "constructive proof" for theorem 1.2 of [Bauer], p. 483. This is a constructive proof because it is based on two explicitly named irrational numbers (√‘2) and (2 logb 9), see sqrt2irrap 12373, 2logb9irrap 15297 and sqrt2cxp2logb9e3 15295. Therefore, this proof is acceptable/usable in intuitionistic logic. (Contributed by Jim Kingdon, 12-Jul-2024.) |
| ⊢ ∃𝑎 ∈ ℝ ∃𝑏 ∈ ℝ (∀𝑝 ∈ ℚ 𝑎 # 𝑝 ∧ ∀𝑞 ∈ ℚ 𝑏 # 𝑞 ∧ (𝑎↑𝑐𝑏) ∈ ℚ) | ||
| Theorem | binom4 15299 | Work out a quartic binomial. (You would think that by this point it would be faster to use binom 11666, but it turns out to be just as much work to put it into this form after clearing all the sums and calculating binomial coefficients.) (Contributed by Mario Carneiro, 6-May-2015.) |
| ⊢ ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → ((𝐴 + 𝐵)↑4) = (((𝐴↑4) + (4 · ((𝐴↑3) · 𝐵))) + ((6 · ((𝐴↑2) · (𝐵↑2))) + ((4 · (𝐴 · (𝐵↑3))) + (𝐵↑4))))) | ||
| Theorem | wilthlem1 15300 | The only elements that are equal to their own inverses in the multiplicative group of nonzero elements in ℤ / 𝑃ℤ are 1 and -1≡𝑃 − 1. (Note that from prmdiveq 12429, (𝑁↑(𝑃 − 2)) mod 𝑃 is the modular inverse of 𝑁 in ℤ / 𝑃ℤ. (Contributed by Mario Carneiro, 24-Jan-2015.) |
| ⊢ ((𝑃 ∈ ℙ ∧ 𝑁 ∈ (1...(𝑃 − 1))) → (𝑁 = ((𝑁↑(𝑃 − 2)) mod 𝑃) ↔ (𝑁 = 1 ∨ 𝑁 = (𝑃 − 1)))) | ||
| < Previous Next > |
| Copyright terms: Public domain | < Previous Next > |