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

Theorem cxpef 25725
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 25724 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
213adant2 1129 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
3 simp2 1135 . . . 4 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → 𝐴 ≠ 0)
43neneqd 2947 . . 3 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → ¬ 𝐴 = 0)
54iffalsed 4467 . 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 1085   = wceq 1539  wcel 2108  wne 2942  ifcif 4456  cfv 6418  (class class class)co 7255  cc 10800  0cc0 10802  1c1 10803   · cmul 10807  expce 15699  logclog 25615  𝑐ccxp 25616
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1799  ax-4 1813  ax-5 1914  ax-6 1972  ax-7 2012  ax-8 2110  ax-9 2118  ax-10 2139  ax-11 2156  ax-12 2173  ax-ext 2709  ax-sep 5218  ax-nul 5225  ax-pr 5347  ax-1cn 10860  ax-icn 10861  ax-addcl 10862  ax-mulcl 10864  ax-i2m1 10870
This theorem depends on definitions:  df-bi 206  df-an 396  df-or 844  df-3an 1087  df-tru 1542  df-fal 1552  df-ex 1784  df-nf 1788  df-sb 2069  df-mo 2540  df-eu 2569  df-clab 2716  df-cleq 2730  df-clel 2817  df-nfc 2888  df-ne 2943  df-ral 3068  df-rex 3069  df-rab 3072  df-v 3424  df-sbc 3712  df-dif 3886  df-un 3888  df-in 3890  df-ss 3900  df-nul 4254  df-if 4457  df-sn 4559  df-pr 4561  df-op 4565  df-uni 4837  df-br 5071  df-opab 5133  df-id 5480  df-xp 5586  df-rel 5587  df-cnv 5588  df-co 5589  df-dm 5590  df-iota 6376  df-fun 6420  df-fv 6426  df-ov 7258  df-oprab 7259  df-mpo 7260  df-cxp 25618
This theorem is referenced by:  cxpexpz  25727  logcxp  25729  1cxp  25732  ecxp  25733  rpcxpcl  25736  cxpne0  25737  cxpadd  25739  mulcxp  25745  cxpmul  25748  abscxp  25752  abscxp2  25753  cxplt  25754  cxple2  25757  cxpsqrtlem  25762  cxpsqrt  25763  cxpefd  25772  1cubrlem  25896  bposlem9  26345  iexpire  33607  aks4d1p1p1  39999
  Copyright terms: Public domain W3C validator