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

Theorem dvds2add 15631
Description: If an integer divides each of two other integers, it divides their sum. (Contributed by Paul Chapman, 21-Mar-2011.)
Assertion
Ref Expression
dvds2add ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾𝑀𝐾𝑁) → 𝐾 ∥ (𝑀 + 𝑁)))

Proof of Theorem dvds2add
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 3simpa 1140 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ))
2 3simpb 1141 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∈ ℤ ∧ 𝑁 ∈ ℤ))
3 zaddcl 12010 . . . 4 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 + 𝑁) ∈ ℤ)
43anim2i 616 . . 3 ((𝐾 ∈ ℤ ∧ (𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ)) → (𝐾 ∈ ℤ ∧ (𝑀 + 𝑁) ∈ ℤ))
543impb 1107 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∈ ℤ ∧ (𝑀 + 𝑁) ∈ ℤ))
6 zaddcl 12010 . . 3 ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) → (𝑥 + 𝑦) ∈ ℤ)
76adantl 482 . 2 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ)) → (𝑥 + 𝑦) ∈ ℤ)
8 zcn 11974 . . . . . . . 8 (𝑥 ∈ ℤ → 𝑥 ∈ ℂ)
9 zcn 11974 . . . . . . . 8 (𝑦 ∈ ℤ → 𝑦 ∈ ℂ)
10 zcn 11974 . . . . . . . 8 (𝐾 ∈ ℤ → 𝐾 ∈ ℂ)
11 adddir 10620 . . . . . . . 8 ((𝑥 ∈ ℂ ∧ 𝑦 ∈ ℂ ∧ 𝐾 ∈ ℂ) → ((𝑥 + 𝑦) · 𝐾) = ((𝑥 · 𝐾) + (𝑦 · 𝐾)))
128, 9, 10, 11syl3an 1152 . . . . . . 7 ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝑥 + 𝑦) · 𝐾) = ((𝑥 · 𝐾) + (𝑦 · 𝐾)))
13123comr 1117 . . . . . 6 ((𝐾 ∈ ℤ ∧ 𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) → ((𝑥 + 𝑦) · 𝐾) = ((𝑥 · 𝐾) + (𝑦 · 𝐾)))
14133expb 1112 . . . . 5 ((𝐾 ∈ ℤ ∧ (𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ)) → ((𝑥 + 𝑦) · 𝐾) = ((𝑥 · 𝐾) + (𝑦 · 𝐾)))
15 oveq12 7154 . . . . 5 (((𝑥 · 𝐾) = 𝑀 ∧ (𝑦 · 𝐾) = 𝑁) → ((𝑥 · 𝐾) + (𝑦 · 𝐾)) = (𝑀 + 𝑁))
1614, 15sylan9eq 2873 . . . 4 (((𝐾 ∈ ℤ ∧ (𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ)) ∧ ((𝑥 · 𝐾) = 𝑀 ∧ (𝑦 · 𝐾) = 𝑁)) → ((𝑥 + 𝑦) · 𝐾) = (𝑀 + 𝑁))
1716ex 413 . . 3 ((𝐾 ∈ ℤ ∧ (𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ)) → (((𝑥 · 𝐾) = 𝑀 ∧ (𝑦 · 𝐾) = 𝑁) → ((𝑥 + 𝑦) · 𝐾) = (𝑀 + 𝑁)))
18173ad2antl1 1177 . 2 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ)) → (((𝑥 · 𝐾) = 𝑀 ∧ (𝑦 · 𝐾) = 𝑁) → ((𝑥 + 𝑦) · 𝐾) = (𝑀 + 𝑁)))
191, 2, 5, 7, 18dvds2lem 15610 1 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾𝑀𝐾𝑁) → 𝐾 ∥ (𝑀 + 𝑁)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 396  w3a 1079   = wceq 1528  wcel 2105   class class class wbr 5057  (class class class)co 7145  cc 10523   + caddc 10528   · cmul 10530  cz 11969  cdvds 15595
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-rab 3144  df-v 3494  df-sbc 3770  df-csb 3881  df-dif 3936  df-un 3938  df-in 3940  df-ss 3949  df-pss 3951  df-nul 4289  df-if 4464  df-pw 4537  df-sn 4558  df-pr 4560  df-tp 4562  df-op 4564  df-uni 4831  df-iun 4912  df-br 5058  df-opab 5120  df-mpt 5138  df-tr 5164  df-id 5453  df-eprel 5458  df-po 5467  df-so 5468  df-fr 5507  df-we 5509  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-pred 6141  df-ord 6187  df-on 6188  df-lim 6189  df-suc 6190  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-om 7570  df-wrecs 7936  df-recs 7997  df-rdg 8035  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-nn 11627  df-n0 11886  df-z 11970  df-dvds 15596
This theorem is referenced by:  dvdssub2  15639  dvdsadd2b  15644  sumeven  15726  divalglem0  15732  sadaddlem  15803  dvdsmulgcd  15893  bezoutr  15900  pythagtriplem19  16158  4sqlem14  16282  4sqlem16  16284  prmdvdsprmop  16367  prmgaplem1  16373  prmgaplcmlem1  16375  dec2dvds  16387  lgsquadlem1  25883  2sqlem8  25929  2sqcoprm  25938  congtr  39440  congadd  39441  jm2.25  39474
  Copyright terms: Public domain W3C validator