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

Theorem replim 14463
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 10626 . 2 (𝐴 ∈ ℂ → ∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)))
2 crre 14461 . . . . . 6 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (ℜ‘(𝑥 + (i · 𝑦))) = 𝑥)
3 crim 14462 . . . . . . 7 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (ℑ‘(𝑥 + (i · 𝑦))) = 𝑦)
43oveq2d 7161 . . . . . 6 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (i · (ℑ‘(𝑥 + (i · 𝑦)))) = (i · 𝑦))
52, 4oveq12d 7163 . . . . 5 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))) = (𝑥 + (i · 𝑦)))
65eqcomd 2824 . . . 4 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (𝑥 + (i · 𝑦)) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))))
7 id 22 . . . . 5 (𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = (𝑥 + (i · 𝑦)))
8 fveq2 6663 . . . . . 6 (𝐴 = (𝑥 + (i · 𝑦)) → (ℜ‘𝐴) = (ℜ‘(𝑥 + (i · 𝑦))))
9 fveq2 6663 . . . . . . 7 (𝐴 = (𝑥 + (i · 𝑦)) → (ℑ‘𝐴) = (ℑ‘(𝑥 + (i · 𝑦))))
109oveq2d 7161 . . . . . 6 (𝐴 = (𝑥 + (i · 𝑦)) → (i · (ℑ‘𝐴)) = (i · (ℑ‘(𝑥 + (i · 𝑦)))))
118, 10oveq12d 7163 . . . . 5 (𝐴 = (𝑥 + (i · 𝑦)) → ((ℜ‘𝐴) + (i · (ℑ‘𝐴))) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦))))))
127, 11eqeq12d 2834 . . . 4 (𝐴 = (𝑥 + (i · 𝑦)) → (𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))) ↔ (𝑥 + (i · 𝑦)) = ((ℜ‘(𝑥 + (i · 𝑦))) + (i · (ℑ‘(𝑥 + (i · 𝑦)))))))
136, 12syl5ibrcom 248 . . 3 ((𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ) → (𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴)))))
1413rexlimivv 3289 . 2 (∃𝑥 ∈ ℝ ∃𝑦 ∈ ℝ 𝐴 = (𝑥 + (i · 𝑦)) → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))
151, 14syl 17 1 (𝐴 ∈ ℂ → 𝐴 = ((ℜ‘𝐴) + (i · (ℑ‘𝐴))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 396   = wceq 1528  wcel 2105  wrex 3136  cfv 6348  (class class class)co 7145  cc 10523  cr 10524  ici 10527   + caddc 10528   · cmul 10530  cre 14444  cim 14445
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1787  ax-4 1801  ax-5 1902  ax-6 1961  ax-7 2006  ax-8 2107  ax-9 2115  ax-10 2136  ax-11 2151  ax-12 2167  ax-ext 2790  ax-sep 5194  ax-nul 5201  ax-pow 5257  ax-pr 5320  ax-un 7450  ax-resscn 10582  ax-1cn 10583  ax-icn 10584  ax-addcl 10585  ax-addrcl 10586  ax-mulcl 10587  ax-mulrcl 10588  ax-mulcom 10589  ax-addass 10590  ax-mulass 10591  ax-distr 10592  ax-i2m1 10593  ax-1ne0 10594  ax-1rid 10595  ax-rnegex 10596  ax-rrecex 10597  ax-cnre 10598  ax-pre-lttri 10599  ax-pre-lttrn 10600  ax-pre-ltadd 10601  ax-pre-mulgt0 10602
This theorem depends on definitions:  df-bi 208  df-an 397  df-or 842  df-3or 1080  df-3an 1081  df-tru 1531  df-ex 1772  df-nf 1776  df-sb 2061  df-mo 2615  df-eu 2647  df-clab 2797  df-cleq 2811  df-clel 2890  df-nfc 2960  df-ne 3014  df-nel 3121  df-ral 3140  df-rex 3141  df-reu 3142  df-rmo 3143  df-rab 3144  df-v 3494  df-sbc 3770  df-csb 3881  df-dif 3936  df-un 3938  df-in 3940  df-ss 3949  df-nul 4289  df-if 4464  df-pw 4537  df-sn 4558  df-pr 4560  df-op 4564  df-uni 4831  df-br 5058  df-opab 5120  df-mpt 5138  df-id 5453  df-po 5467  df-so 5468  df-xp 5554  df-rel 5555  df-cnv 5556  df-co 5557  df-dm 5558  df-rn 5559  df-res 5560  df-ima 5561  df-iota 6307  df-fun 6350  df-fn 6351  df-f 6352  df-f1 6353  df-fo 6354  df-f1o 6355  df-fv 6356  df-riota 7103  df-ov 7148  df-oprab 7149  df-mpo 7150  df-er 8278  df-en 8498  df-dom 8499  df-sdom 8500  df-pnf 10665  df-mnf 10666  df-xr 10667  df-ltxr 10668  df-le 10669  df-sub 10860  df-neg 10861  df-div 11286  df-2 11688  df-cj 14446  df-re 14447  df-im 14448
This theorem is referenced by:  remim  14464  reim0b  14466  rereb  14467  mulre  14468  cjreb  14470  reneg  14472  readd  14473  remullem  14475  imneg  14480  imadd  14481  cjcj  14487  imval2  14498  cnrecnv  14512  replimi  14517  replimd  14544  recan  14684  efeul  15503  absef  15538  absefib  15539  efieq1re  15540  cnsubrg  20533  itgconst  24346  tanregt0  25050  tanarg  25129  ccfldextdgrr  30956
  Copyright terms: Public domain W3C validator