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

Theorem cxpef 26707
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 26706 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
213adant2 1132 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))))
3 simp2 1138 . . . 4 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → 𝐴 ≠ 0)
43neneqd 2945 . . 3 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → ¬ 𝐴 = 0)
54iffalsed 4536 . 2 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → if(𝐴 = 0, if(𝐵 = 0, 1, 0), (exp‘(𝐵 · (log‘𝐴)))) = (exp‘(𝐵 · (log‘𝐴))))
62, 5eqtrd 2777 1 ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐵 ∈ ℂ) → (𝐴𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  w3a 1087   = wceq 1540  wcel 2108  wne 2940  ifcif 4525  cfv 6561  (class class class)co 7431  cc 11153  0cc0 11155  1c1 11156   · cmul 11160  expce 16097  logclog 26596  𝑐ccxp 26597
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1795  ax-4 1809  ax-5 1910  ax-6 1967  ax-7 2007  ax-8 2110  ax-9 2118  ax-10 2141  ax-11 2157  ax-12 2177  ax-ext 2708  ax-sep 5296  ax-nul 5306  ax-pr 5432  ax-1cn 11213  ax-icn 11214  ax-addcl 11215  ax-mulcl 11217  ax-i2m1 11223
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 849  df-3an 1089  df-tru 1543  df-fal 1553  df-ex 1780  df-nf 1784  df-sb 2065  df-mo 2540  df-eu 2569  df-clab 2715  df-cleq 2729  df-clel 2816  df-nfc 2892  df-ne 2941  df-ral 3062  df-rex 3071  df-rab 3437  df-v 3482  df-sbc 3789  df-dif 3954  df-un 3956  df-ss 3968  df-nul 4334  df-if 4526  df-sn 4627  df-pr 4629  df-op 4633  df-uni 4908  df-br 5144  df-opab 5206  df-id 5578  df-xp 5691  df-rel 5692  df-cnv 5693  df-co 5694  df-dm 5695  df-iota 6514  df-fun 6563  df-fv 6569  df-ov 7434  df-oprab 7435  df-mpo 7436  df-cxp 26599
This theorem is referenced by:  cxpexpz  26709  logcxp  26711  1cxp  26714  ecxp  26715  rpcxpcl  26718  cxpne0  26719  cxpadd  26721  mulcxp  26727  cxpmul  26730  abscxp  26734  abscxp2  26735  cxplt  26736  cxple2  26739  cxpsqrtlem  26744  cxpsqrt  26745  cxpefd  26754  1cubrlem  26884  bposlem9  27336  iexpire  35735  aks4d1p1p1  42064
  Copyright terms: Public domain W3C validator