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

Theorem gcddvds 16384
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 12511 . . . . . 6 0 ∈ ℤ
2 dvds0 16155 . . . . . 6 (0 ∈ ℤ → 0 ∥ 0)
31, 2ax-mp 5 . . . . 5 0 ∥ 0
4 breq2 5110 . . . . . . 7 (𝑀 = 0 → (0 ∥ 𝑀 ↔ 0 ∥ 0))
5 breq2 5110 . . . . . . 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 7367 . . . . . . 7 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = (0 gcd 0))
11 gcd0val 16378 . . . . . . 7 (0 gcd 0) = 0
1210, 11eqtrdi 2793 . . . . . 6 ((𝑀 = 0 ∧ 𝑁 = 0) → (𝑀 gcd 𝑁) = 0)
1312breq1d 5116 . . . . 5 ((𝑀 = 0 ∧ 𝑁 = 0) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ 0 ∥ 𝑀))
1412breq1d 5116 . . . . 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 2737 . . . . 5 {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧} = {𝑛 ∈ ℤ ∣ ∀𝑧 ∈ {𝑀, 𝑁}𝑛𝑧}
19 eqid 2737 . . . . 5 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}
2018, 19gcdcllem3 16382 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∈ ℕ ∧ (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁) ∧ ((𝐾 ∈ ℤ ∧ 𝐾𝑀𝐾𝑁) → 𝐾 ≤ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))))
2120simp2d 1144 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀 ∧ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑁))
22 gcdn0val 16379 . . . . 5 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → (𝑀 gcd 𝑁) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ))
2322breq1d 5116 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∧ 𝑁 = 0)) → ((𝑀 gcd 𝑁) ∥ 𝑀 ↔ sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) ∥ 𝑀))
2422breq1d 5116 . . . 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 3065  {crab 3408  {cpr 4589   class class class wbr 5106  (class class class)co 7358  supcsup 9377  cr 11051  0cc0 11052   < clt 11190  cle 11191  cn 12154  cz 12500  cdvds 16137   gcd cgcd 16375
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 2708  ax-sep 5257  ax-nul 5264  ax-pow 5321  ax-pr 5385  ax-un 7673  ax-cnex 11108  ax-resscn 11109  ax-1cn 11110  ax-icn 11111  ax-addcl 11112  ax-addrcl 11113  ax-mulcl 11114  ax-mulrcl 11115  ax-mulcom 11116  ax-addass 11117  ax-mulass 11118  ax-distr 11119  ax-i2m1 11120  ax-1ne0 11121  ax-1rid 11122  ax-rnegex 11123  ax-rrecex 11124  ax-cnre 11125  ax-pre-lttri 11126  ax-pre-lttrn 11127  ax-pre-ltadd 11128  ax-pre-mulgt0 11129  ax-pre-sup 11130
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 2539  df-eu 2568  df-clab 2715  df-cleq 2729  df-clel 2815  df-nfc 2890  df-ne 2945  df-nel 3051  df-ral 3066  df-rex 3075  df-rmo 3354  df-reu 3355  df-rab 3409  df-v 3448  df-sbc 3741  df-csb 3857  df-dif 3914  df-un 3916  df-in 3918  df-ss 3928  df-pss 3930  df-nul 4284  df-if 4488  df-pw 4563  df-sn 4588  df-pr 4590  df-op 4594  df-uni 4867  df-iun 4957  df-br 5107  df-opab 5169  df-mpt 5190  df-tr 5224  df-id 5532  df-eprel 5538  df-po 5546  df-so 5547  df-fr 5589  df-we 5591  df-xp 5640  df-rel 5641  df-cnv 5642  df-co 5643  df-dm 5644  df-rn 5645  df-res 5646  df-ima 5647  df-pred 6254  df-ord 6321  df-on 6322  df-lim 6323  df-suc 6324  df-iota 6449  df-fun 6499  df-fn 6500  df-f 6501  df-f1 6502  df-fo 6503  df-f1o 6504  df-fv 6505  df-riota 7314  df-ov 7361  df-oprab 7362  df-mpo 7363  df-om 7804  df-2nd 7923  df-frecs 8213  df-wrecs 8244  df-recs 8318  df-rdg 8357  df-er 8649  df-en 8885  df-dom 8886  df-sdom 8887  df-sup 9379  df-inf 9380  df-pnf 11192  df-mnf 11193  df-xr 11194  df-ltxr 11195  df-le 11196  df-sub 11388  df-neg 11389  df-div 11814  df-nn 12155  df-2 12217  df-3 12218  df-n0 12415  df-z 12501  df-uz 12765  df-rp 12917  df-seq 13908  df-exp 13969  df-cj 14985  df-re 14986  df-im 14987  df-sqrt 15121  df-abs 15122  df-dvds 16138  df-gcd 16376
This theorem is referenced by:  zeqzmulgcd  16391  divgcdz  16392  divgcdnn  16396  gcd0id  16400  gcdneg  16403  gcdaddmlem  16405  gcd1  16409  bezoutlem4  16424  dvdsgcdb  16427  dfgcd2  16428  mulgcd  16430  gcdzeq  16434  dvdsmulgcd  16437  sqgcd  16442  dvdssqlem  16443  bezoutr  16445  gcddvdslcm  16479  lcmgcdlem  16483  lcmgcdeq  16489  coprmgcdb  16526  mulgcddvds  16532  rpmulgcd2  16533  qredeu  16535  rpdvds  16537  divgcdcoprm0  16542  divgcdodd  16587  coprm  16588  rpexp  16599  divnumden  16624  phimullem  16652  hashgcdlem  16661  hashgcdeq  16662  phisum  16663  pythagtriplem4  16692  pythagtriplem19  16706  pcgcd1  16750  pc2dvds  16752  pockthlem  16778  odmulg  19339  odadd1  19627  odadd2  19628  znunit  20973  znrrg  20975  dvdsmulf1o  26546  2sqlem8  26777  2sqcoprm  26786  qqhval2lem  32565  aks4d1p8d1  40544  gcdle1d  40819  gcdle2d  40820  expgcd  40823  dvdsexpnn  40829  fltdvdsabdvdsc  40979  goldbachthlem2  45745  divgcdoddALTV  45881
  Copyright terms: Public domain W3C validator