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

Theorem cxpef 26692
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 26691 . . 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 2935 . . 3 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → ¬ 𝐴 = 0)
54iffalsed 4544 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))) = (exp‘(𝐵 · (log‘𝐴))))
62, 5eqtrd 2766 1 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  w3a 1084   = wceq 1534  wcel 2099  wne 2930  ifcif 4533  cfv 6554  (class class class)co 7424  cc 11156  0cc0 11158  1c1 11159   · cmul 11163  expce 16063  logclog 26581  𝑐ccxp 26582
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1790  ax-4 1804  ax-5 1906  ax-6 1964  ax-7 2004  ax-8 2101  ax-9 2109  ax-10 2130  ax-11 2147  ax-12 2167  ax-ext 2697  ax-sep 5304  ax-nul 5311  ax-pr 5433  ax-1cn 11216  ax-icn 11217  ax-addcl 11218  ax-mulcl 11220  ax-i2m1 11226
This theorem depends on definitions:  df-bi 206  df-an 395  df-or 846  df-3an 1086  df-tru 1537  df-fal 1547  df-ex 1775  df-nf 1779  df-sb 2061  df-mo 2529  df-eu 2558  df-clab 2704  df-cleq 2718  df-clel 2803  df-nfc 2878  df-ne 2931  df-ral 3052  df-rex 3061  df-rab 3420  df-v 3464  df-sbc 3777  df-dif 3950  df-un 3952  df-ss 3964  df-nul 4326  df-if 4534  df-sn 4634  df-pr 4636  df-op 4640  df-uni 4914  df-br 5154  df-opab 5216  df-id 5580  df-xp 5688  df-rel 5689  df-cnv 5690  df-co 5691  df-dm 5692  df-iota 6506  df-fun 6556  df-fv 6562  df-ov 7427  df-oprab 7428  df-mpo 7429  df-cxp 26584
This theorem is referenced by:  cxpexpz  26694  logcxp  26696  1cxp  26699  ecxp  26700  rpcxpcl  26703  cxpne0  26704  cxpadd  26706  mulcxp  26712  cxpmul  26715  abscxp  26719  abscxp2  26720  cxplt  26721  cxple2  26724  cxpsqrtlem  26729  cxpsqrt  26730  cxpefd  26739  1cubrlem  26869  bposlem9  27321  iexpire  35557  aks4d1p1p1  41762
  Copyright terms: Public domain W3C validator