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

Theorem replim 14827
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 10972 . 2 (𝐴 ∈ ℂ → ∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)))
2 crre 14825 . . . . . 6 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (ℜ‘(𝑥 + (i · 𝑦))) = 𝑥)
3 crim 14826 . . . . . . 7 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (ℑ‘(𝑥 + (i · 𝑦))) = 𝑦)
43oveq2d 7291 . . . . . 6 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (i · (ℑ‘(𝑥 + (i · 𝑦)))) = (i · 𝑦))
52, 4oveq12d 7293 . . . . 5 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))) = (𝑥 + (i · 𝑦)))
65eqcomd 2744 . . . 4 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (𝑥 + (i · 𝑦)) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))))
7 id 22 . . . . 5 (𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = (𝑥 + (i · 𝑦)))
8 fveq2 6774 . . . . . 6 (𝐴 = (𝑥 + (i · 𝑦)) → (ℜ‘𝐴) = (ℜ‘(𝑥 + (i · 𝑦))))
9 fveq2 6774 . . . . . . 7 (𝐴 = (𝑥 + (i · 𝑦)) → (ℑ‘𝐴) = (ℑ‘(𝑥 + (i · 𝑦))))
109oveq2d 7291 . . . . . 6 (𝐴 = (𝑥 + (i · 𝑦)) → (i · (ℑ‘𝐴)) = (i · (ℑ‘(𝑥 + (i · 𝑦)))))
118, 10oveq12d 7293 . . . . 5 (𝐴 = (𝑥 + (i · 𝑦)) → ((ℜ‘𝐴) + (i · (ℑ‘𝐴))) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))))
127, 11eqeq12d 2754 . . . 4 (𝐴 = (𝑥 + (i · 𝑦)) → (𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))) ↔ (𝑥 + (i · 𝑦)) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦)))))))
136, 12syl5ibrcom 246 . . 3 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴)))))
1413rexlimivv 3221 . 2 (∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))
151, 14syl 17 1 (𝐴 ∈ ℂ → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 396   = wceq 1539  wcel 2106  wrex 3065  cfv 6433  (class class class)co 7275  cc 10869  cr 10870  ici 10873   + caddc 10874   · cmul 10876  cre 14808  cim 14809
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1798  ax-4 1812  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-10 2137  ax-11 2154  ax-12 2171  ax-ext 2709  ax-sep 5223  ax-nul 5230  ax-pow 5288  ax-pr 5352  ax-un 7588  ax-resscn 10928  ax-1cn 10929  ax-icn 10930  ax-addcl 10931  ax-addrcl 10932  ax-mulcl 10933  ax-mulrcl 10934  ax-mulcom 10935  ax-addass 10936  ax-mulass 10937  ax-distr 10938  ax-i2m1 10939  ax-1ne0 10940  ax-1rid 10941  ax-rnegex 10942  ax-rrecex 10943  ax-cnre 10944  ax-pre-lttri 10945  ax-pre-lttrn 10946  ax-pre-ltadd 10947  ax-pre-mulgt0 10948
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 845  df-3or 1087  df-3an 1088  df-tru 1542  df-fal 1552  df-ex 1783  df-nf 1787  df-sb 2068  df-mo 2540  df-eu 2569  df-clab 2716  df-cleq 2730  df-clel 2816  df-nfc 2889  df-ne 2944  df-nel 3050  df-ral 3069  df-rex 3070  df-rmo 3071  df-reu 3072  df-rab 3073  df-v 3434  df-sbc 3717  df-csb 3833  df-dif 3890  df-un 3892  df-in 3894  df-ss 3904  df-nul 4257  df-if 4460  df-pw 4535  df-sn 4562  df-pr 4564  df-op 4568  df-uni 4840  df-br 5075  df-opab 5137  df-mpt 5158  df-id 5489  df-po 5503  df-so 5504  df-xp 5595  df-rel 5596  df-cnv 5597  df-co 5598  df-dm 5599  df-rn 5600  df-res 5601  df-ima 5602  df-iota 6391  df-fun 6435  df-fn 6436  df-f 6437  df-f1 6438  df-fo 6439  df-f1o 6440  df-fv 6441  df-riota 7232  df-ov 7278  df-oprab 7279  df-mpo 7280  df-er 8498  df-en 8734  df-dom 8735  df-sdom 8736  df-pnf 11011  df-mnf 11012  df-xr 11013  df-ltxr 11014  df-le 11015  df-sub 11207  df-neg 11208  df-div 11633  df-2 12036  df-cj 14810  df-re 14811  df-im 14812
This theorem is referenced by:  remim  14828  reim0b  14830  rereb  14831  mulre  14832  cjreb  14834  reneg  14836  readd  14837  remullem  14839  imneg  14844  imadd  14845  cjcj  14851  imval2  14862  cnrecnv  14876  replimi  14881  replimd  14908  recan  15048  efeul  15871  absef  15906  absefib  15907  efieq1re  15908  cnsubrg  20658  itgconst  24983  tanregt0  25695  tanarg  25774  ccfldextdgrr  31742  sqrtcval  41249
  Copyright terms: Public domain W3C validator