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

Theorem gcddvds 16191
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 12313 . . . . . 6 0 ∈ ℤ
2 dvds0 15962 . . . . . 6 (0 ∈ ℤ → 0 ∥ 0)
31, 2ax-mp 5 . . . . 5 0 ∥ 0
4 breq2 5082 . . . . . . 7 (𝑀 = 0 → (0 ∥ 𝑀 ↔ 0 ∥ 0))
5 breq2 5082 . . . . . . 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 7277 . . . . . . 7 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = (0 gcd 0))
11 gcd0val 16185 . . . . . . 7 (0 gcd 0) = 0
1210, 11eqtrdi 2795 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = 0)
1312breq1d 5088 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ 0 ∥ 𝑀))
1412breq1d 5088 . . . . 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 2739 . . . . 5 {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧} = {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧}
19 eqid 2739 . . . . 5 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}
2018, 19gcdcllem3 16189 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∈ ℕ ∧ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁) ∧ ((𝐾 ∈ ℤ ∧ 𝐾𝑀𝐾𝑁) → 𝐾 ≤ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))))
2120simp2d 1141 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
22 gcdn0val 16186 . . . . 5 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (𝑀 gcd 𝑁) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))
2322breq1d 5088 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀))
2422breq1d 5088 . . . 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 1541  wcel 2109  wral 3065  {crab 3069  {cpr 4568   class class class wbr 5078  (class class class)co 7268  supcsup 9160  cr 10854  0cc0 10855   < clt 10993  cle 10994  cn 11956  cz 12302  cdvds 15944   gcd cgcd 16182
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1801  ax-4 1815  ax-5 1916  ax-6 1974  ax-7 2014  ax-8 2111  ax-9 2119  ax-10 2140  ax-11 2157  ax-12 2174  ax-ext 2710  ax-sep 5226  ax-nul 5233  ax-pow 5291  ax-pr 5355  ax-un 7579  ax-cnex 10911  ax-resscn 10912  ax-1cn 10913  ax-icn 10914  ax-addcl 10915  ax-addrcl 10916  ax-mulcl 10917  ax-mulrcl 10918  ax-mulcom 10919  ax-addass 10920  ax-mulass 10921  ax-distr 10922  ax-i2m1 10923  ax-1ne0 10924  ax-1rid 10925  ax-rnegex 10926  ax-rrecex 10927  ax-cnre 10928  ax-pre-lttri 10929  ax-pre-lttrn 10930  ax-pre-ltadd 10931  ax-pre-mulgt0 10932  ax-pre-sup 10933
This theorem depends on definitions:  df-bi 206  df-an 396  df-or 844  df-3or 1086  df-3an 1087  df-tru 1544  df-fal 1554  df-ex 1786  df-nf 1790  df-sb 2071  df-mo 2541  df-eu 2570  df-clab 2717  df-cleq 2731  df-clel 2817  df-nfc 2890  df-ne 2945  df-nel 3051  df-ral 3070  df-rex 3071  df-reu 3072  df-rmo 3073  df-rab 3074  df-v 3432  df-sbc 3720  df-csb 3837  df-dif 3894  df-un 3896  df-in 3898  df-ss 3908  df-pss 3910  df-nul 4262  df-if 4465  df-pw 4540  df-sn 4567  df-pr 4569  df-tp 4571  df-op 4573  df-uni 4845  df-iun 4931  df-br 5079  df-opab 5141  df-mpt 5162  df-tr 5196  df-id 5488  df-eprel 5494  df-po 5502  df-so 5503  df-fr 5543  df-we 5545  df-xp 5594  df-rel 5595  df-cnv 5596  df-co 5597  df-dm 5598  df-rn 5599  df-res 5600  df-ima 5601  df-pred 6199  df-ord 6266  df-on 6267  df-lim 6268  df-suc 6269  df-iota 6388  df-fun 6432  df-fn 6433  df-f 6434  df-f1 6435  df-fo 6436  df-f1o 6437  df-fv 6438  df-riota 7225  df-ov 7271  df-oprab 7272  df-mpo 7273  df-om 7701  df-2nd 7818  df-frecs 8081  df-wrecs 8112  df-recs 8186  df-rdg 8225  df-er 8472  df-en 8708  df-dom 8709  df-sdom 8710  df-sup 9162  df-inf 9163  df-pnf 10995  df-mnf 10996  df-xr 10997  df-ltxr 10998  df-le 10999  df-sub 11190  df-neg 11191  df-div 11616  df-nn 11957  df-2 12019  df-3 12020  df-n0 12217  df-z 12303  df-uz 12565  df-rp 12713  df-seq 13703  df-exp 13764  df-cj 14791  df-re 14792  df-im 14793  df-sqrt 14927  df-abs 14928  df-dvds 15945  df-gcd 16183
This theorem is referenced by:  zeqzmulgcd  16198  divgcdz  16199  divgcdnn  16203  gcd0id  16207  gcdneg  16210  gcdaddmlem  16212  gcd1  16216  bezoutlem4  16231  dvdsgcdb  16234  dfgcd2  16235  mulgcd  16237  gcdzeq  16243  dvdsmulgcd  16246  sqgcd  16251  dvdssqlem  16252  bezoutr  16254  gcddvdslcm  16288  lcmgcdlem  16292  lcmgcdeq  16298  coprmgcdb  16335  mulgcddvds  16341  rpmulgcd2  16342  qredeu  16344  rpdvds  16346  divgcdcoprm0  16351  divgcdodd  16396  coprm  16397  rpexp  16408  divnumden  16433  phimullem  16461  hashgcdlem  16470  hashgcdeq  16471  phisum  16472  pythagtriplem4  16501  pythagtriplem19  16515  pcgcd1  16559  pc2dvds  16561  pockthlem  16587  odmulg  19144  odadd1  19430  odadd2  19431  znunit  20752  znrrg  20754  dvdsmulf1o  26324  2sqlem8  26555  2sqcoprm  26564  qqhval2lem  31910  aks4d1p8d1  40072  gcdle1d  40310  gcdle2d  40311  expgcd  40314  dvdsexpnn  40320  fltdvdsabdvdsc  40455  goldbachthlem2  44950  divgcdoddALTV  45086
  Copyright terms: Public domain W3C validator