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

Theorem gcddvds 16219
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 12339 . . . . . 6 0 ∈ ℤ
2 dvds0 15990 . . . . . 6 (0 ∈ ℤ → 0 ∥ 0)
31, 2ax-mp 5 . . . . 5 0 ∥ 0
4 breq2 5079 . . . . . . 7 (𝑀 = 0 → (0 ∥ 𝑀 ↔ 0 ∥ 0))
5 breq2 5079 . . . . . . 7 (𝑁 = 0 → (0 ∥ 𝑁 ↔ 0 ∥ 0))
64, 5bi2anan9 636 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → ((0 ∥ 𝑀 ∧ 0 ∥ 𝑁) ↔ (0 ∥ 0 ∧ 0 ∥ 0)))
7 anidm 565 . . . . . 6 ((0 ∥ 0 ∧ 0 ∥ 0) ↔ 0 ∥ 0)
86, 7bitrdi 287 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((0 ∥ 𝑀 ∧ 0 ∥ 𝑁) ↔ 0 ∥ 0))
93, 8mpbiri 257 . . . 4 ((𝑀 = 0 ∧ 𝑁 = 0) → (0 ∥ 𝑀 ∧ 0 ∥ 𝑁))
10 oveq12 7293 . . . . . . 7 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = (0 gcd 0))
11 gcd0val 16213 . . . . . . 7 (0 gcd 0) = 0
1210, 11eqtrdi 2795 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = 0)
1312breq1d 5085 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ 0 ∥ 𝑀))
1412breq1d 5085 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑁 ↔ 0 ∥ 𝑁))
1513, 14anbi12d 631 . . . 4 ((𝑀 = 0 ∧ 𝑁 = 0) → (((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁) ↔ (0 ∥ 𝑀 ∧ 0 ∥ 𝑁)))
169, 15mpbird 256 . . 3 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
1716adantl 482 . 2 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
18 eqid 2739 . . . . 5 {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧} = {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧}
19 eqid 2739 . . . . 5 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}
2018, 19gcdcllem3 16217 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∈ ℕ ∧ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁) ∧ ((𝐾 ∈ ℤ ∧ 𝐾𝑀𝐾𝑁) → 𝐾 ≤ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))))
2120simp2d 1142 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
22 gcdn0val 16214 . . . . 5 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (𝑀 gcd 𝑁) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))
2322breq1d 5085 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀))
2422breq1d 5085 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑁 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
2523, 24anbi12d 631 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁) ↔ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁)))
2621, 25mpbird 256 . 2 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
2717, 26pm2.61dan 810 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 gcd 𝑁) ∥ 𝑀 ∧ (𝑀 gcd 𝑁) ∥ 𝑁))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 396  w3a 1086   = wceq 1539  wcel 2107  wral 3065  {crab 3069  {cpr 4564   class class class wbr 5075  (class class class)co 7284  supcsup 9208  cr 10879  0cc0 10880   < clt 11018  cle 11019  cn 11982  cz 12328  cdvds 15972   gcd cgcd 16210
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 2710  ax-sep 5224  ax-nul 5231  ax-pow 5289  ax-pr 5353  ax-un 7597  ax-cnex 10936  ax-resscn 10937  ax-1cn 10938  ax-icn 10939  ax-addcl 10940  ax-addrcl 10941  ax-mulcl 10942  ax-mulrcl 10943  ax-mulcom 10944  ax-addass 10945  ax-mulass 10946  ax-distr 10947  ax-i2m1 10948  ax-1ne0 10949  ax-1rid 10950  ax-rnegex 10951  ax-rrecex 10952  ax-cnre 10953  ax-pre-lttri 10954  ax-pre-lttrn 10955  ax-pre-ltadd 10956  ax-pre-mulgt0 10957  ax-pre-sup 10958
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 845  df-3or 1087  df-3an 1088  df-tru 1542  df-fal 1552  df-ex 1783  df-nf 1787  df-sb 2069  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-rmo 3072  df-reu 3073  df-rab 3074  df-v 3435  df-sbc 3718  df-csb 3834  df-dif 3891  df-un 3893  df-in 3895  df-ss 3905  df-pss 3907  df-nul 4258  df-if 4461  df-pw 4536  df-sn 4563  df-pr 4565  df-op 4569  df-uni 4841  df-iun 4927  df-br 5076  df-opab 5138  df-mpt 5159  df-tr 5193  df-id 5490  df-eprel 5496  df-po 5504  df-so 5505  df-fr 5545  df-we 5547  df-xp 5596  df-rel 5597  df-cnv 5598  df-co 5599  df-dm 5600  df-rn 5601  df-res 5602  df-ima 5603  df-pred 6206  df-ord 6273  df-on 6274  df-lim 6275  df-suc 6276  df-iota 6395  df-fun 6439  df-fn 6440  df-f 6441  df-f1 6442  df-fo 6443  df-f1o 6444  df-fv 6445  df-riota 7241  df-ov 7287  df-oprab 7288  df-mpo 7289  df-om 7722  df-2nd 7841  df-frecs 8106  df-wrecs 8137  df-recs 8211  df-rdg 8250  df-er 8507  df-en 8743  df-dom 8744  df-sdom 8745  df-sup 9210  df-inf 9211  df-pnf 11020  df-mnf 11021  df-xr 11022  df-ltxr 11023  df-le 11024  df-sub 11216  df-neg 11217  df-div 11642  df-nn 11983  df-2 12045  df-3 12046  df-n0 12243  df-z 12329  df-uz 12592  df-rp 12740  df-seq 13731  df-exp 13792  df-cj 14819  df-re 14820  df-im 14821  df-sqrt 14955  df-abs 14956  df-dvds 15973  df-gcd 16211
This theorem is referenced by:  zeqzmulgcd  16226  divgcdz  16227  divgcdnn  16231  gcd0id  16235  gcdneg  16238  gcdaddmlem  16240  gcd1  16244  bezoutlem4  16259  dvdsgcdb  16262  dfgcd2  16263  mulgcd  16265  gcdzeq  16271  dvdsmulgcd  16274  sqgcd  16279  dvdssqlem  16280  bezoutr  16282  gcddvdslcm  16316  lcmgcdlem  16320  lcmgcdeq  16326  coprmgcdb  16363  mulgcddvds  16369  rpmulgcd2  16370  qredeu  16372  rpdvds  16374  divgcdcoprm0  16379  divgcdodd  16424  coprm  16425  rpexp  16436  divnumden  16461  phimullem  16489  hashgcdlem  16498  hashgcdeq  16499  phisum  16500  pythagtriplem4  16529  pythagtriplem19  16543  pcgcd1  16587  pc2dvds  16589  pockthlem  16615  odmulg  19172  odadd1  19458  odadd2  19459  znunit  20780  znrrg  20782  dvdsmulf1o  26352  2sqlem8  26583  2sqcoprm  26592  qqhval2lem  31940  aks4d1p8d1  40099  gcdle1d  40337  gcdle2d  40338  expgcd  40341  dvdsexpnn  40347  fltdvdsabdvdsc  40482  goldbachthlem2  45009  divgcdoddALTV  45145
  Copyright terms: Public domain W3C validator