MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  cxpef Structured version   Visualization version   GIF version

Theorem cxpef 25820
Description: Value of the complex power function. (Contributed by Mario Carneiro, 2-Aug-2014.)
Assertion
Ref Expression
cxpef ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴))))

Proof of Theorem cxpef
StepHypRef Expression
1 cxpval 25819 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
213adant2 1130 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
3 simp2 1136 . . . 4 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → 𝐴 ≠ 0)
43neneqd 2948 . . 3 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → ¬ 𝐴 = 0)
54iffalsed 4470 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))) = (exp‘(𝐵 · (log‘𝐴))))
62, 5eqtrd 2778 1 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  w3a 1086   = wceq 1539  wcel 2106  wne 2943  ifcif 4459  cfv 6433  (class class class)co 7275  cc 10869  0cc0 10871  1c1 10872   · cmul 10876  expce 15771  logclog 25710  𝑐ccxp 25711
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1798  ax-4 1812  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-10 2137  ax-11 2154  ax-12 2171  ax-ext 2709  ax-sep 5223  ax-nul 5230  ax-pr 5352  ax-1cn 10929  ax-icn 10930  ax-addcl 10931  ax-mulcl 10933  ax-i2m1 10939
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 845  df-3an 1088  df-tru 1542  df-fal 1552  df-ex 1783  df-nf 1787  df-sb 2068  df-mo 2540  df-eu 2569  df-clab 2716  df-cleq 2730  df-clel 2816  df-nfc 2889  df-ne 2944  df-ral 3069  df-rex 3070  df-rab 3073  df-v 3434  df-sbc 3717  df-dif 3890  df-un 3892  df-in 3894  df-ss 3904  df-nul 4257  df-if 4460  df-sn 4562  df-pr 4564  df-op 4568  df-uni 4840  df-br 5075  df-opab 5137  df-id 5489  df-xp 5595  df-rel 5596  df-cnv 5597  df-co 5598  df-dm 5599  df-iota 6391  df-fun 6435  df-fv 6441  df-ov 7278  df-oprab 7279  df-mpo 7280  df-cxp 25713
This theorem is referenced by:  cxpexpz  25822  logcxp  25824  1cxp  25827  ecxp  25828  rpcxpcl  25831  cxpne0  25832  cxpadd  25834  mulcxp  25840  cxpmul  25843  abscxp  25847  abscxp2  25848  cxplt  25849  cxple2  25852  cxpsqrtlem  25857  cxpsqrt  25858  cxpefd  25867  1cubrlem  25991  bposlem9  26440  iexpire  33701  aks4d1p1p1  40071
  Copyright terms: Public domain W3C validator