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

Theorem gcddvds 16138
Description: The gcd of two integers divides each of them. (Contributed by Paul Chapman, 21-Mar-2011.)
Assertion
Ref Expression
gcddvds ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))

Proof of Theorem gcddvds
Dummy variables 𝑛 𝐾 𝑧 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 0z 12260 . . . . . 6 0 ∈ ℤ
2 dvds0 15909 . . . . . 6 (0 ∈ ℤ → 0 ∥ 0)
31, 2ax-mp 5 . . . . 5 0 ∥ 0
4 breq2 5074 . . . . . . 7 (𝑀 = 0 → (0 ∥ 𝑀 ↔ 0 ∥ 0))
5 breq2 5074 . . . . . . 7 (𝑁 = 0 → (0 ∥ 𝑁 ↔ 0 ∥ 0))
64, 5bi2anan9 635 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → ((0 ∥ 𝑀 ∧ 0 ∥ 𝑁) ↔ (0 ∥ 0 ∧ 0 ∥ 0)))
7 anidm 564 . . . . . 6 ((0 ∥ 0 ∧ 0 ∥ 0) ↔ 0 ∥ 0)
86, 7bitrdi 286 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((0 ∥ 𝑀 ∧ 0 ∥ 𝑁) ↔ 0 ∥ 0))
93, 8mpbiri 257 . . . 4 ((𝑀 = 0 ∧ 𝑁 = 0) → (0 ∥ 𝑀 ∧ 0 ∥ 𝑁))
10 oveq12 7264 . . . . . . 7 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = (0 gcd 0))
11 gcd0val 16132 . . . . . . 7 (0 gcd 0) = 0
1210, 11eqtrdi 2795 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = 0)
1312breq1d 5080 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ 0 ∥ 𝑀))
1412breq1d 5080 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑁 ↔ 0 ∥ 𝑁))
1513, 14anbi12d 630 . . . 4 ((𝑀 = 0 ∧ 𝑁 = 0) → (((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁) ↔ (0 ∥ 𝑀 ∧ 0 ∥ 𝑁)))
169, 15mpbird 256 . . 3 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
1716adantl 481 . 2 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
18 eqid 2738 . . . . 5 {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧} = {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧}
19 eqid 2738 . . . . 5 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}
2018, 19gcdcllem3 16136 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∈ ℕ ∧ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁) ∧ ((𝐾 ∈ ℤ ∧ 𝐾𝑀𝐾𝑁) → 𝐾 ≤ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))))
2120simp2d 1141 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
22 gcdn0val 16133 . . . . 5 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (𝑀 gcd 𝑁) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))
2322breq1d 5080 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀))
2422breq1d 5080 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑁 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
2523, 24anbi12d 630 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁) ↔ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁)))
2621, 25mpbird 256 . 2 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
2717, 26pm2.61dan 809 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 395  w3a 1085   = wceq 1539  wcel 2108  wral 3063  {crab 3067  {cpr 4560   class class class wbr 5070  (class class class)co 7255  supcsup 9129  cr 10801  0cc0 10802   < clt 10940  cle 10941  cn 11903  cz 12249  cdvds 15891   gcd cgcd 16129
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1799  ax-4 1813  ax-5 1914  ax-6 1972  ax-7 2012  ax-8 2110  ax-9 2118  ax-10 2139  ax-11 2156  ax-12 2173  ax-ext 2709  ax-sep 5218  ax-nul 5225  ax-pow 5283  ax-pr 5347  ax-un 7566  ax-cnex 10858  ax-resscn 10859  ax-1cn 10860  ax-icn 10861  ax-addcl 10862  ax-addrcl 10863  ax-mulcl 10864  ax-mulrcl 10865  ax-mulcom 10866  ax-addass 10867  ax-mulass 10868  ax-distr 10869  ax-i2m1 10870  ax-1ne0 10871  ax-1rid 10872  ax-rnegex 10873  ax-rrecex 10874  ax-cnre 10875  ax-pre-lttri 10876  ax-pre-lttrn 10877  ax-pre-ltadd 10878  ax-pre-mulgt0 10879  ax-pre-sup 10880
This theorem depends on definitions:  df-bi 206  df-an 396  df-or 844  df-3or 1086  df-3an 1087  df-tru 1542  df-fal 1552  df-ex 1784  df-nf 1788  df-sb 2069  df-mo 2540  df-eu 2569  df-clab 2716  df-cleq 2730  df-clel 2817  df-nfc 2888  df-ne 2943  df-nel 3049  df-ral 3068  df-rex 3069  df-reu 3070  df-rmo 3071  df-rab 3072  df-v 3424  df-sbc 3712  df-csb 3829  df-dif 3886  df-un 3888  df-in 3890  df-ss 3900  df-pss 3902  df-nul 4254  df-if 4457  df-pw 4532  df-sn 4559  df-pr 4561  df-tp 4563  df-op 4565  df-uni 4837  df-iun 4923  df-br 5071  df-opab 5133  df-mpt 5154  df-tr 5188  df-id 5480  df-eprel 5486  df-po 5494  df-so 5495  df-fr 5535  df-we 5537  df-xp 5586  df-rel 5587  df-cnv 5588  df-co 5589  df-dm 5590  df-rn 5591  df-res 5592  df-ima 5593  df-pred 6191  df-ord 6254  df-on 6255  df-lim 6256  df-suc 6257  df-iota 6376  df-fun 6420  df-fn 6421  df-f 6422  df-f1 6423  df-fo 6424  df-f1o 6425  df-fv 6426  df-riota 7212  df-ov 7258  df-oprab 7259  df-mpo 7260  df-om 7688  df-2nd 7805  df-frecs 8068  df-wrecs 8099  df-recs 8173  df-rdg 8212  df-er 8456  df-en 8692  df-dom 8693  df-sdom 8694  df-sup 9131  df-inf 9132  df-pnf 10942  df-mnf 10943  df-xr 10944  df-ltxr 10945  df-le 10946  df-sub 11137  df-neg 11138  df-div 11563  df-nn 11904  df-2 11966  df-3 11967  df-n0 12164  df-z 12250  df-uz 12512  df-rp 12660  df-seq 13650  df-exp 13711  df-cj 14738  df-re 14739  df-im 14740  df-sqrt 14874  df-abs 14875  df-dvds 15892  df-gcd 16130
This theorem is referenced by:  zeqzmulgcd  16145  divgcdz  16146  divgcdnn  16150  gcd0id  16154  gcdneg  16157  gcdaddmlem  16159  gcd1  16163  bezoutlem4  16178  dvdsgcdb  16181  dfgcd2  16182  mulgcd  16184  gcdzeq  16190  dvdsmulgcd  16193  sqgcd  16198  dvdssqlem  16199  bezoutr  16201  gcddvdslcm  16235  lcmgcdlem  16239  lcmgcdeq  16245  coprmgcdb  16282  mulgcddvds  16288  rpmulgcd2  16289  qredeu  16291  rpdvds  16293  divgcdcoprm0  16298  divgcdodd  16343  coprm  16344  rpexp  16355  divnumden  16380  phimullem  16408  hashgcdlem  16417  hashgcdeq  16418  phisum  16419  pythagtriplem4  16448  pythagtriplem19  16462  pcgcd1  16506  pc2dvds  16508  pockthlem  16534  odmulg  19078  odadd1  19364  odadd2  19365  znunit  20683  znrrg  20685  dvdsmulf1o  26248  2sqlem8  26479  2sqcoprm  26488  qqhval2lem  31831  aks4d1p8d1  40020  gcdle1d  40251  gcdle2d  40252  expgcd  40255  dvdsexpnn  40261  fltdvdsabdvdsc  40391  goldbachthlem2  44886  divgcdoddALTV  45022
  Copyright terms: Public domain W3C validator