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

Theorem gcdcom 16530
Description: The gcd operator is commutative. Theorem 1.4(a) in [ApostolNT] p. 16. (Contributed by Paul Chapman, 21-Mar-2011.)
Assertion
Ref Expression
gcdcom ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = (𝑁 gcd 𝑀))

Proof of Theorem gcdcom
Dummy variable 𝑛 is distinct from all other variables.
StepHypRef Expression
1 ancom 460 . . 3 ((𝑀 = 0 ∧ 𝑁 = 0) ↔ (𝑁 = 0 ∧ 𝑀 = 0))
2 ancom 460 . . . . 5 ((𝑛𝑀𝑛𝑁) ↔ (𝑛𝑁𝑛𝑀))
32rabbii 3421 . . . 4 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}
43supeq1i 9457 . . 3 sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )
51, 4ifbieq2i 4526 . 2 if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < ))
6 gcdval 16513 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )))
7 gcdval 16513 . . 3 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
87ancoms 458 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
95, 6, 83eqtr4a 2796 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = (𝑁 gcd 𝑀))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 395   = wceq 1540  wcel 2108  {crab 3415  ifcif 4500   class class class wbr 5119  (class class class)co 7403  supcsup 9450  cr 11126  0cc0 11127   < clt 11267  cz 12586  cdvds 16270   gcd cgcd 16511
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1795  ax-4 1809  ax-5 1910  ax-6 1967  ax-7 2007  ax-8 2110  ax-9 2118  ax-10 2141  ax-11 2157  ax-12 2177  ax-ext 2707  ax-sep 5266  ax-nul 5276  ax-pow 5335  ax-pr 5402  ax-un 7727  ax-resscn 11184  ax-1cn 11185  ax-icn 11186  ax-addcl 11187  ax-mulcl 11189  ax-i2m1 11195  ax-pre-lttri 11201  ax-pre-lttrn 11202
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 848  df-3or 1087  df-3an 1088  df-tru 1543  df-fal 1553  df-ex 1780  df-nf 1784  df-sb 2065  df-mo 2539  df-eu 2568  df-clab 2714  df-cleq 2727  df-clel 2809  df-nfc 2885  df-ne 2933  df-nel 3037  df-ral 3052  df-rex 3061  df-rmo 3359  df-rab 3416  df-v 3461  df-sbc 3766  df-csb 3875  df-dif 3929  df-un 3931  df-in 3933  df-ss 3943  df-nul 4309  df-if 4501  df-pw 4577  df-sn 4602  df-pr 4604  df-op 4608  df-uni 4884  df-br 5120  df-opab 5182  df-mpt 5202  df-id 5548  df-po 5561  df-so 5562  df-xp 5660  df-rel 5661  df-cnv 5662  df-co 5663  df-dm 5664  df-rn 5665  df-res 5666  df-ima 5667  df-iota 6483  df-fun 6532  df-fn 6533  df-f 6534  df-f1 6535  df-fo 6536  df-f1o 6537  df-fv 6538  df-ov 7406  df-oprab 7407  df-mpo 7408  df-er 8717  df-en 8958  df-dom 8959  df-sdom 8960  df-sup 9452  df-pnf 11269  df-mnf 11270  df-ltxr 11272  df-gcd 16512
This theorem is referenced by:  gcdcomd  16531  divgcdnnr  16533  gcdid0  16537  neggcd  16540  gcdabs2  16547  1gcd  16550  6gcd4e2  16555  rprpwr  16576  eucalginv  16601  3lcm2e6woprm  16632  coprmdvds  16670  qredeq  16674  divgcdcoprmex  16683  cncongr1  16684  cncongrprm  16746  fermltl  16801  vfermltl  16819  coprimeprodsq2  16827  pythagtrip  16852  pcgcd  16896  pockthlem  16923  gcdi  17091  gcdmodi  17092  1259lem5  17152  2503lem3  17156  4001lem4  17161  odinv  19540  lgsprme0  27300  lgsdirnn0  27305  lgsquad2lem2  27346  lgsquad3  27348  ex-gcd  30384  gcd32  35712  gcdcomnni  41947  aks6d1c1  42075  aks6d1c4  42083  goldbachthlem2  47508  goldbachth  47509  gcd2odd1  47630  fpprwpprb  47702
  Copyright terms: Public domain W3C validator