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

Theorem gcddvds 16441
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 12566 . . . . . 6 0 ∈ ℤ
2 dvds0 16212 . . . . . 6 (0 ∈ ℤ → 0 ∥ 0)
31, 2ax-mp 5 . . . . 5 0 ∥ 0
4 breq2 5152 . . . . . . 7 (𝑀 = 0 → (0 ∥ 𝑀 ↔ 0 ∥ 0))
5 breq2 5152 . . . . . . 7 (𝑁 = 0 → (0 ∥ 𝑁 ↔ 0 ∥ 0))
64, 5bi2anan9 638 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → ((0 ∥ 𝑀 ∧ 0 ∥ 𝑁) ↔ (0 ∥ 0 ∧ 0 ∥ 0)))
7 anidm 566 . . . . . 6 ((0 ∥ 0 ∧ 0 ∥ 0) ↔ 0 ∥ 0)
86, 7bitrdi 287 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((0 ∥ 𝑀 ∧ 0 ∥ 𝑁) ↔ 0 ∥ 0))
93, 8mpbiri 258 . . . 4 ((𝑀 = 0 ∧ 𝑁 = 0) → (0 ∥ 𝑀 ∧ 0 ∥ 𝑁))
10 oveq12 7415 . . . . . . 7 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = (0 gcd 0))
11 gcd0val 16435 . . . . . . 7 (0 gcd 0) = 0
1210, 11eqtrdi 2789 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = 0)
1312breq1d 5158 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ 0 ∥ 𝑀))
1412breq1d 5158 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑁 ↔ 0 ∥ 𝑁))
1513, 14anbi12d 632 . . . 4 ((𝑀 = 0 ∧ 𝑁 = 0) → (((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁) ↔ (0 ∥ 𝑀 ∧ 0 ∥ 𝑁)))
169, 15mpbird 257 . . 3 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
1716adantl 483 . 2 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
18 eqid 2733 . . . . 5 {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧} = {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧}
19 eqid 2733 . . . . 5 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}
2018, 19gcdcllem3 16439 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∈ ℕ ∧ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁) ∧ ((𝐾 ∈ ℤ ∧ 𝐾𝑀𝐾𝑁) → 𝐾 ≤ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))))
2120simp2d 1144 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
22 gcdn0val 16436 . . . . 5 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (𝑀 gcd 𝑁) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))
2322breq1d 5158 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀))
2422breq1d 5158 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑁 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
2523, 24anbi12d 632 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁) ↔ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁)))
2621, 25mpbird 257 . 2 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
2717, 26pm2.61dan 812 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 397  w3a 1088   = wceq 1542  wcel 2107  wral 3062  {crab 3433  {cpr 4630   class class class wbr 5148  (class class class)co 7406  supcsup 9432  cr 11106  0cc0 11107   < clt 11245  cle 11246  cn 12209  cz 12555  cdvds 16194   gcd cgcd 16432
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 1914  ax-6 1972  ax-7 2012  ax-8 2109  ax-9 2117  ax-10 2138  ax-11 2155  ax-12 2172  ax-ext 2704  ax-sep 5299  ax-nul 5306  ax-pow 5363  ax-pr 5427  ax-un 7722  ax-cnex 11163  ax-resscn 11164  ax-1cn 11165  ax-icn 11166  ax-addcl 11167  ax-addrcl 11168  ax-mulcl 11169  ax-mulrcl 11170  ax-mulcom 11171  ax-addass 11172  ax-mulass 11173  ax-distr 11174  ax-i2m1 11175  ax-1ne0 11176  ax-1rid 11177  ax-rnegex 11178  ax-rrecex 11179  ax-cnre 11180  ax-pre-lttri 11181  ax-pre-lttrn 11182  ax-pre-ltadd 11183  ax-pre-mulgt0 11184  ax-pre-sup 11185
This theorem depends on definitions:  df-bi 206  df-an 398  df-or 847  df-3or 1089  df-3an 1090  df-tru 1545  df-fal 1555  df-ex 1783  df-nf 1787  df-sb 2069  df-mo 2535  df-eu 2564  df-clab 2711  df-cleq 2725  df-clel 2811  df-nfc 2886  df-ne 2942  df-nel 3048  df-ral 3063  df-rex 3072  df-rmo 3377  df-reu 3378  df-rab 3434  df-v 3477  df-sbc 3778  df-csb 3894  df-dif 3951  df-un 3953  df-in 3955  df-ss 3965  df-pss 3967  df-nul 4323  df-if 4529  df-pw 4604  df-sn 4629  df-pr 4631  df-op 4635  df-uni 4909  df-iun 4999  df-br 5149  df-opab 5211  df-mpt 5232  df-tr 5266  df-id 5574  df-eprel 5580  df-po 5588  df-so 5589  df-fr 5631  df-we 5633  df-xp 5682  df-rel 5683  df-cnv 5684  df-co 5685  df-dm 5686  df-rn 5687  df-res 5688  df-ima 5689  df-pred 6298  df-ord 6365  df-on 6366  df-lim 6367  df-suc 6368  df-iota 6493  df-fun 6543  df-fn 6544  df-f 6545  df-f1 6546  df-fo 6547  df-f1o 6548  df-fv 6549  df-riota 7362  df-ov 7409  df-oprab 7410  df-mpo 7411  df-om 7853  df-2nd 7973  df-frecs 8263  df-wrecs 8294  df-recs 8368  df-rdg 8407  df-er 8700  df-en 8937  df-dom 8938  df-sdom 8939  df-sup 9434  df-inf 9435  df-pnf 11247  df-mnf 11248  df-xr 11249  df-ltxr 11250  df-le 11251  df-sub 11443  df-neg 11444  df-div 11869  df-nn 12210  df-2 12272  df-3 12273  df-n0 12470  df-z 12556  df-uz 12820  df-rp 12972  df-seq 13964  df-exp 14025  df-cj 15043  df-re 15044  df-im 15045  df-sqrt 15179  df-abs 15180  df-dvds 16195  df-gcd 16433
This theorem is referenced by:  zeqzmulgcd  16448  divgcdz  16449  divgcdnn  16453  gcd0id  16457  gcdneg  16460  gcdaddmlem  16462  gcd1  16466  bezoutlem4  16481  dvdsgcdb  16484  dfgcd2  16485  mulgcd  16487  gcdzeq  16491  dvdsmulgcd  16494  sqgcd  16499  dvdssqlem  16500  bezoutr  16502  gcddvdslcm  16536  lcmgcdlem  16540  lcmgcdeq  16546  coprmgcdb  16583  mulgcddvds  16589  rpmulgcd2  16590  qredeu  16592  rpdvds  16594  divgcdcoprm0  16599  divgcdodd  16644  coprm  16645  rpexp  16656  divnumden  16681  phimullem  16709  hashgcdlem  16718  hashgcdeq  16719  phisum  16720  pythagtriplem4  16749  pythagtriplem19  16763  pcgcd1  16807  pc2dvds  16809  pockthlem  16835  odmulg  19419  odadd1  19711  odadd2  19712  znunit  21111  znrrg  21113  dvdsmulf1o  26688  2sqlem8  26919  2sqcoprm  26928  qqhval2lem  32950  aks4d1p8d1  40938  gcdle1d  41217  gcdle2d  41218  expgcd  41221  dvdsexpnn  41227  fltdvdsabdvdsc  41377  goldbachthlem2  46201  divgcdoddALTV  46337
  Copyright terms: Public domain W3C validator