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

Theorem divdir 10710
Description: Distribution of division over addition. (Contributed by NM, 31-Jul-2004.) (Proof shortened by Mario Carneiro, 27-May-2016.)
Assertion
Ref Expression
divdir ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → ((𝐴 + 𝐵) / 𝐶) = ((𝐴 / 𝐶) + (𝐵 / 𝐶)))

Proof of Theorem divdir
StepHypRef Expression
1 simp1 1061 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → 𝐴 ∈ ℂ)
2 simp2 1062 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → 𝐵 ∈ ℂ)
3 reccl 10692 . . . 4 ((𝐶 ∈ ℂ ∧ 𝐶 ≠ 0) → (1 / 𝐶) ∈ ℂ)
433ad2ant3 1084 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → (1 / 𝐶) ∈ ℂ)
51, 2, 4adddird 10065 . 2 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → ((𝐴 + 𝐵) · (1 / 𝐶)) = ((𝐴 · (1 / 𝐶)) + (𝐵 · (1 / 𝐶))))
61, 2addcld 10059 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → (𝐴 + 𝐵) ∈ ℂ)
7 simp3l 1089 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → 𝐶 ∈ ℂ)
8 simp3r 1090 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → 𝐶 ≠ 0)
9 divrec 10701 . . 3 (((𝐴 + 𝐵) ∈ ℂ ∧ 𝐶 ∈ ℂ ∧ 𝐶 ≠ 0) → ((𝐴 + 𝐵) / 𝐶) = ((𝐴 + 𝐵) · (1 / 𝐶)))
106, 7, 8, 9syl3anc 1326 . 2 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → ((𝐴 + 𝐵) / 𝐶) = ((𝐴 + 𝐵) · (1 / 𝐶)))
11 divrec 10701 . . . 4 ((𝐴 ∈ ℂ ∧ 𝐶 ∈ ℂ ∧ 𝐶 ≠ 0) → (𝐴 / 𝐶) = (𝐴 · (1 / 𝐶)))
121, 7, 8, 11syl3anc 1326 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → (𝐴 / 𝐶) = (𝐴 · (1 / 𝐶)))
13 divrec 10701 . . . 4 ((𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ ∧ 𝐶 ≠ 0) → (𝐵 / 𝐶) = (𝐵 · (1 / 𝐶)))
142, 7, 8, 13syl3anc 1326 . . 3 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → (𝐵 / 𝐶) = (𝐵 · (1 / 𝐶)))
1512, 14oveq12d 6668 . 2 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → ((𝐴 / 𝐶) + (𝐵 / 𝐶)) = ((𝐴 · (1 / 𝐶)) + (𝐵 · (1 / 𝐶))))
165, 10, 153eqtr4d 2666 1 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ (𝐶 ∈ ℂ ∧ 𝐶 ≠ 0)) → ((𝐴 + 𝐵) / 𝐶) = ((𝐴 / 𝐶) + (𝐵 / 𝐶)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 384  w3a 1037   = wceq 1483  wcel 1990  wne 2794  (class class class)co 6650  cc 9934  0cc0 9936  1c1 9937   + caddc 9939   · cmul 9941   / cdiv 10684
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1722  ax-4 1737  ax-5 1839  ax-6 1888  ax-7 1935  ax-8 1992  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602  ax-sep 4781  ax-nul 4789  ax-pow 4843  ax-pr 4906  ax-un 6949  ax-resscn 9993  ax-1cn 9994  ax-icn 9995  ax-addcl 9996  ax-addrcl 9997  ax-mulcl 9998  ax-mulrcl 9999  ax-mulcom 10000  ax-addass 10001  ax-mulass 10002  ax-distr 10003  ax-i2m1 10004  ax-1ne0 10005  ax-1rid 10006  ax-rnegex 10007  ax-rrecex 10008  ax-cnre 10009  ax-pre-lttri 10010  ax-pre-lttrn 10011  ax-pre-ltadd 10012  ax-pre-mulgt0 10013
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-3or 1038  df-3an 1039  df-tru 1486  df-ex 1705  df-nf 1710  df-sb 1881  df-eu 2474  df-mo 2475  df-clab 2609  df-cleq 2615  df-clel 2618  df-nfc 2753  df-ne 2795  df-nel 2898  df-ral 2917  df-rex 2918  df-reu 2919  df-rmo 2920  df-rab 2921  df-v 3202  df-sbc 3436  df-csb 3534  df-dif 3577  df-un 3579  df-in 3581  df-ss 3588  df-nul 3916  df-if 4087  df-pw 4160  df-sn 4178  df-pr 4180  df-op 4184  df-uni 4437  df-br 4654  df-opab 4713  df-mpt 4730  df-id 5024  df-po 5035  df-so 5036  df-xp 5120  df-rel 5121  df-cnv 5122  df-co 5123  df-dm 5124  df-rn 5125  df-res 5126  df-ima 5127  df-iota 5851  df-fun 5890  df-fn 5891  df-f 5892  df-f1 5893  df-fo 5894  df-f1o 5895  df-fv 5896  df-riota 6611  df-ov 6653  df-oprab 6654  df-mpt2 6655  df-er 7742  df-en 7956  df-dom 7957  df-sdom 7958  df-pnf 10076  df-mnf 10077  df-xr 10078  df-ltxr 10079  df-le 10080  df-sub 10268  df-neg 10269  df-div 10685
This theorem is referenced by:  muldivdir  10720  divsubdir  10721  divadddiv  10740  divdirzi  10777  divdird  10839  2halves  11260  halfaddsub  11265  zdivadd  11448  nneo  11461  rpnnen1lem5  11818  rpnnen1lem5OLD  11824  2tnp1ge0ge0  12630  fldiv  12659  modcyc  12705  mulsubdivbinom2  13046  crim  13855  efival  14882  flodddiv4  15137  divgcdcoprm0  15379  pythagtriplem17  15536  ptolemy  24248  relogbmul  24515  harmonicbnd4  24737  ppiub  24929  logfacrlim  24949  bposlem9  25017  2lgslem3a  25121  2lgslem3b  25122  2lgslem3c  25123  2lgslem3d  25124  chpchtlim  25168  mudivsum  25219  selberglem2  25235  pntrsumo1  25254  pntibndlem2  25280  pntibndlem3  25281  pntlemb  25286  dpfrac1  29599  dpfrac1OLD  29600  heiborlem6  33615  zofldiv2ALTV  41574  zofldiv2  42325  sinhpcosh  42481  onetansqsecsq  42502
  Copyright terms: Public domain W3C validator