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

Theorem gcdcom 16491
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 459 . . 3 ((𝑀 = 0 ∧ 𝑁 = 0) ↔ (𝑁 = 0 ∧ 𝑀 = 0))
2 ancom 459 . . . . 5 ((𝑛𝑀𝑛𝑁) ↔ (𝑛𝑁𝑛𝑀))
32rabbii 3424 . . . 4 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}
43supeq1i 9472 . . 3 sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )
51, 4ifbieq2i 4555 . 2 if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < ))
6 gcdval 16474 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )))
7 gcdval 16474 . . 3 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
87ancoms 457 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
95, 6, 83eqtr4a 2791 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = (𝑁 gcd 𝑀))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 394   = wceq 1533  wcel 2098  {crab 3418  ifcif 4530   class class class wbr 5149  (class class class)co 7419  supcsup 9465  cr 11139  0cc0 11140   < clt 11280  cz 12591  cdvds 16234   gcd cgcd 16472
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1789  ax-4 1803  ax-5 1905  ax-6 1963  ax-7 2003  ax-8 2100  ax-9 2108  ax-10 2129  ax-11 2146  ax-12 2166  ax-ext 2696  ax-sep 5300  ax-nul 5307  ax-pow 5365  ax-pr 5429  ax-un 7741  ax-resscn 11197  ax-1cn 11198  ax-icn 11199  ax-addcl 11200  ax-mulcl 11202  ax-i2m1 11208  ax-pre-lttri 11214  ax-pre-lttrn 11215
This theorem depends on definitions:  df-bi 206  df-an 395  df-or 846  df-3or 1085  df-3an 1086  df-tru 1536  df-fal 1546  df-ex 1774  df-nf 1778  df-sb 2060  df-mo 2528  df-eu 2557  df-clab 2703  df-cleq 2717  df-clel 2802  df-nfc 2877  df-ne 2930  df-nel 3036  df-ral 3051  df-rex 3060  df-rmo 3363  df-rab 3419  df-v 3463  df-sbc 3774  df-csb 3890  df-dif 3947  df-un 3949  df-in 3951  df-ss 3961  df-nul 4323  df-if 4531  df-pw 4606  df-sn 4631  df-pr 4633  df-op 4637  df-uni 4910  df-br 5150  df-opab 5212  df-mpt 5233  df-id 5576  df-po 5590  df-so 5591  df-xp 5684  df-rel 5685  df-cnv 5686  df-co 5687  df-dm 5688  df-rn 5689  df-res 5690  df-ima 5691  df-iota 6501  df-fun 6551  df-fn 6552  df-f 6553  df-f1 6554  df-fo 6555  df-f1o 6556  df-fv 6557  df-ov 7422  df-oprab 7423  df-mpo 7424  df-er 8725  df-en 8965  df-dom 8966  df-sdom 8967  df-sup 9467  df-pnf 11282  df-mnf 11283  df-ltxr 11285  df-gcd 16473
This theorem is referenced by:  gcdcomd  16492  divgcdnnr  16494  gcdid0  16498  neggcd  16501  gcdabs2  16508  1gcd  16512  6gcd4e2  16517  rprpwr  16537  eucalginv  16558  3lcm2e6woprm  16589  coprmdvds  16627  qredeq  16631  divgcdcoprmex  16640  cncongr1  16641  cncongrprm  16704  fermltl  16756  vfermltl  16773  coprimeprodsq2  16781  pythagtrip  16806  pcgcd  16850  pockthlem  16877  gcdi  17045  gcdmodi  17046  1259lem5  17107  2503lem3  17111  4001lem4  17116  odinv  19528  lgsprme0  27317  lgsdirnn0  27322  lgsquad2lem2  27363  lgsquad3  27365  ex-gcd  30339  gcd32  35474  gcdcomnni  41591  aks6d1c1  41719  aks6d1c4  41727  goldbachthlem2  47023  goldbachth  47024  gcd2odd1  47145  fpprwpprb  47217
  Copyright terms: Public domain W3C validator