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

Theorem cxpef 25256
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 25255 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
213adant2 1128 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
3 simp2 1134 . . . 4 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → 𝐴 ≠ 0)
43neneqd 2992 . . 3 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → ¬ 𝐴 = 0)
54iffalsed 4436 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))) = (exp‘(𝐵 · (log‘𝐴))))
62, 5eqtrd 2833 1 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  w3a 1084   = wceq 1538  wcel 2111  wne 2987  ifcif 4425  cfv 6324  (class class class)co 7135  cc 10524  0cc0 10526  1c1 10527   · cmul 10531  expce 15407  logclog 25146  𝑐ccxp 25147
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1797  ax-4 1811  ax-5 1911  ax-6 1970  ax-7 2015  ax-8 2113  ax-9 2121  ax-10 2142  ax-11 2158  ax-12 2175  ax-ext 2770  ax-sep 5167  ax-nul 5174  ax-pr 5295  ax-1cn 10584  ax-icn 10585  ax-addcl 10586  ax-mulcl 10588  ax-i2m1 10594
This theorem depends on definitions:  df-bi 210  df-an 400  df-or 845  df-3an 1086  df-tru 1541  df-ex 1782  df-nf 1786  df-sb 2070  df-mo 2598  df-eu 2629  df-clab 2777  df-cleq 2791  df-clel 2870  df-nfc 2938  df-ne 2988  df-ral 3111  df-rex 3112  df-rab 3115  df-v 3443  df-sbc 3721  df-dif 3884  df-un 3886  df-in 3888  df-ss 3898  df-nul 4244  df-if 4426  df-sn 4526  df-pr 4528  df-op 4532  df-uni 4801  df-br 5031  df-opab 5093  df-id 5425  df-xp 5525  df-rel 5526  df-cnv 5527  df-co 5528  df-dm 5529  df-iota 6283  df-fun 6326  df-fv 6332  df-ov 7138  df-oprab 7139  df-mpo 7140  df-cxp 25149
This theorem is referenced by:  cxpexpz  25258  logcxp  25260  1cxp  25263  ecxp  25264  rpcxpcl  25267  cxpne0  25268  cxpadd  25270  mulcxp  25276  cxpmul  25279  abscxp  25283  abscxp2  25284  cxplt  25285  cxple2  25288  cxpsqrtlem  25293  cxpsqrt  25294  cxpefd  25303  1cubrlem  25427  bposlem9  25876  iexpire  33080  aks4d1p1p1  39345
  Copyright terms: Public domain W3C validator