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

Theorem gcdcom 16546
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 3438 . . . 4 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}
43supeq1i 9484 . . 3 sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )
51, 4ifbieq2i 4555 . 2 if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < ))
6 gcdval 16529 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )))
7 gcdval 16529 . . 3 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
87ancoms 458 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
95, 6, 83eqtr4a 2800 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = (𝑁 gcd 𝑀))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 395   = wceq 1536  wcel 2105  {crab 3432  ifcif 4530   class class class wbr 5147  (class class class)co 7430  supcsup 9477  cr 11151  0cc0 11152   < clt 11292  cz 12610  cdvds 16286   gcd cgcd 16527
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1791  ax-4 1805  ax-5 1907  ax-6 1964  ax-7 2004  ax-8 2107  ax-9 2115  ax-10 2138  ax-11 2154  ax-12 2174  ax-ext 2705  ax-sep 5301  ax-nul 5311  ax-pow 5370  ax-pr 5437  ax-un 7753  ax-resscn 11209  ax-1cn 11210  ax-icn 11211  ax-addcl 11212  ax-mulcl 11214  ax-i2m1 11220  ax-pre-lttri 11226  ax-pre-lttrn 11227
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 848  df-3or 1087  df-3an 1088  df-tru 1539  df-fal 1549  df-ex 1776  df-nf 1780  df-sb 2062  df-mo 2537  df-eu 2566  df-clab 2712  df-cleq 2726  df-clel 2813  df-nfc 2889  df-ne 2938  df-nel 3044  df-ral 3059  df-rex 3068  df-rmo 3377  df-rab 3433  df-v 3479  df-sbc 3791  df-csb 3908  df-dif 3965  df-un 3967  df-in 3969  df-ss 3979  df-nul 4339  df-if 4531  df-pw 4606  df-sn 4631  df-pr 4633  df-op 4637  df-uni 4912  df-br 5148  df-opab 5210  df-mpt 5231  df-id 5582  df-po 5596  df-so 5597  df-xp 5694  df-rel 5695  df-cnv 5696  df-co 5697  df-dm 5698  df-rn 5699  df-res 5700  df-ima 5701  df-iota 6515  df-fun 6564  df-fn 6565  df-f 6566  df-f1 6567  df-fo 6568  df-f1o 6569  df-fv 6570  df-ov 7433  df-oprab 7434  df-mpo 7435  df-er 8743  df-en 8984  df-dom 8985  df-sdom 8986  df-sup 9479  df-pnf 11294  df-mnf 11295  df-ltxr 11297  df-gcd 16528
This theorem is referenced by:  gcdcomd  16547  divgcdnnr  16549  gcdid0  16553  neggcd  16556  gcdabs2  16563  1gcd  16566  6gcd4e2  16571  rprpwr  16592  eucalginv  16617  3lcm2e6woprm  16648  coprmdvds  16686  qredeq  16690  divgcdcoprmex  16699  cncongr1  16700  cncongrprm  16762  fermltl  16817  vfermltl  16834  coprimeprodsq2  16842  pythagtrip  16867  pcgcd  16911  pockthlem  16938  gcdi  17106  gcdmodi  17107  1259lem5  17168  2503lem3  17172  4001lem4  17177  odinv  19593  lgsprme0  27397  lgsdirnn0  27402  lgsquad2lem2  27443  lgsquad3  27445  ex-gcd  30485  gcd32  35728  gcdcomnni  41969  aks6d1c1  42097  aks6d1c4  42105  goldbachthlem2  47470  goldbachth  47471  gcd2odd1  47592  fpprwpprb  47664
  Copyright terms: Public domain W3C validator