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

Theorem replim 15155
Description: Reconstruct a complex number from its real and imaginary parts. (Contributed by NM, 10-May-1999.) (Revised by Mario Carneiro, 7-Nov-2013.)
Assertion
Ref Expression
replim (𝐴 ∈ ℂ → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))

Proof of Theorem replim
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 cnre 11258 . 2 (𝐴 ∈ ℂ → ∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)))
2 crre 15153 . . . . . 6 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (ℜ‘(𝑥 + (i · 𝑦))) = 𝑥)
3 crim 15154 . . . . . . 7 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (ℑ‘(𝑥 + (i · 𝑦))) = 𝑦)
43oveq2d 7447 . . . . . 6 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (i · (ℑ‘(𝑥 + (i · 𝑦)))) = (i · 𝑦))
52, 4oveq12d 7449 . . . . 5 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))) = (𝑥 + (i · 𝑦)))
65eqcomd 2743 . . . 4 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (𝑥 + (i · 𝑦)) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))))
7 id 22 . . . . 5 (𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = (𝑥 + (i · 𝑦)))
8 fveq2 6906 . . . . . 6 (𝐴 = (𝑥 + (i · 𝑦)) → (ℜ‘𝐴) = (ℜ‘(𝑥 + (i · 𝑦))))
9 fveq2 6906 . . . . . . 7 (𝐴 = (𝑥 + (i · 𝑦)) → (ℑ‘𝐴) = (ℑ‘(𝑥 + (i · 𝑦))))
109oveq2d 7447 . . . . . 6 (𝐴 = (𝑥 + (i · 𝑦)) → (i · (ℑ‘𝐴)) = (i · (ℑ‘(𝑥 + (i · 𝑦)))))
118, 10oveq12d 7449 . . . . 5 (𝐴 = (𝑥 + (i · 𝑦)) → ((ℜ‘𝐴) + (i · (ℑ‘𝐴))) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))))
127, 11eqeq12d 2753 . . . 4 (𝐴 = (𝑥 + (i · 𝑦)) → (𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))) ↔ (𝑥 + (i · 𝑦)) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦)))))))
136, 12syl5ibrcom 247 . . 3 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴)))))
1413rexlimivv 3201 . 2 (∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))
151, 14syl 17 1 (𝐴 ∈ ℂ → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 395   = wceq 1540  wcel 2108  wrex 3070  cfv 6561  (class class class)co 7431  cc 11153  cr 11154  ici 11157   + caddc 11158   · cmul 11160  cre 15136  cim 15137
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-pow 5365  ax-pr 5432  ax-un 7755  ax-resscn 11212  ax-1cn 11213  ax-icn 11214  ax-addcl 11215  ax-addrcl 11216  ax-mulcl 11217  ax-mulrcl 11218  ax-mulcom 11219  ax-addass 11220  ax-mulass 11221  ax-distr 11222  ax-i2m1 11223  ax-1ne0 11224  ax-1rid 11225  ax-rnegex 11226  ax-rrecex 11227  ax-cnre 11228  ax-pre-lttri 11229  ax-pre-lttrn 11230  ax-pre-ltadd 11231  ax-pre-mulgt0 11232
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 849  df-3or 1088  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-nel 3047  df-ral 3062  df-rex 3071  df-rmo 3380  df-reu 3381  df-rab 3437  df-v 3482  df-sbc 3789  df-csb 3900  df-dif 3954  df-un 3956  df-in 3958  df-ss 3968  df-nul 4334  df-if 4526  df-pw 4602  df-sn 4627  df-pr 4629  df-op 4633  df-uni 4908  df-br 5144  df-opab 5206  df-mpt 5226  df-id 5578  df-po 5592  df-so 5593  df-xp 5691  df-rel 5692  df-cnv 5693  df-co 5694  df-dm 5695  df-rn 5696  df-res 5697  df-ima 5698  df-iota 6514  df-fun 6563  df-fn 6564  df-f 6565  df-f1 6566  df-fo 6567  df-f1o 6568  df-fv 6569  df-riota 7388  df-ov 7434  df-oprab 7435  df-mpo 7436  df-er 8745  df-en 8986  df-dom 8987  df-sdom 8988  df-pnf 11297  df-mnf 11298  df-xr 11299  df-ltxr 11300  df-le 11301  df-sub 11494  df-neg 11495  df-div 11921  df-2 12329  df-cj 15138  df-re 15139  df-im 15140
This theorem is referenced by:  remim  15156  reim0b  15158  rereb  15159  mulre  15160  cjreb  15162  reneg  15164  readd  15165  remullem  15167  imneg  15172  imadd  15173  cjcj  15179  imval2  15190  cnrecnv  15204  replimi  15209  replimd  15236  recan  15375  efeul  16198  absef  16233  absefib  16234  efieq1re  16235  cnsubrg  21445  itgconst  25854  tanregt0  26581  tanarg  26661  ccfldextdgrr  33722  sqrtcval  43654
  Copyright terms: Public domain W3C validator