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

Theorem gcdcom 16559
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 3449 . . . 4 {𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)} = {𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}
43supeq1i 9516 . . 3 sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < ) = sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )
51, 4ifbieq2i 4573 . 2 if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < ))
6 gcdval 16542 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = if((𝑀 = 0 ∧ 𝑁 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑀𝑛𝑁)}, ℝ, < )))
7 gcdval 16542 . . 3 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
87ancoms 458 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 gcd 𝑀) = if((𝑁 = 0 ∧ 𝑀 = 0), 0, sup({𝑛 ∈ ℤ ∣ (𝑛𝑁𝑛𝑀)}, ℝ, < )))
95, 6, 83eqtr4a 2806 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 gcd 𝑁) = (𝑁 gcd 𝑀))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 395   = wceq 1537  wcel 2108  {crab 3443  ifcif 4548   class class class wbr 5166  (class class class)co 7448  supcsup 9509  cr 11183  0cc0 11184   < clt 11324  cz 12639  cdvds 16302   gcd cgcd 16540
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1793  ax-4 1807  ax-5 1909  ax-6 1967  ax-7 2007  ax-8 2110  ax-9 2118  ax-10 2141  ax-11 2158  ax-12 2178  ax-ext 2711  ax-sep 5317  ax-nul 5324  ax-pow 5383  ax-pr 5447  ax-un 7770  ax-resscn 11241  ax-1cn 11242  ax-icn 11243  ax-addcl 11244  ax-mulcl 11246  ax-i2m1 11252  ax-pre-lttri 11258  ax-pre-lttrn 11259
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 847  df-3or 1088  df-3an 1089  df-tru 1540  df-fal 1550  df-ex 1778  df-nf 1782  df-sb 2065  df-mo 2543  df-eu 2572  df-clab 2718  df-cleq 2732  df-clel 2819  df-nfc 2895  df-ne 2947  df-nel 3053  df-ral 3068  df-rex 3077  df-rmo 3388  df-rab 3444  df-v 3490  df-sbc 3805  df-csb 3922  df-dif 3979  df-un 3981  df-in 3983  df-ss 3993  df-nul 4353  df-if 4549  df-pw 4624  df-sn 4649  df-pr 4651  df-op 4655  df-uni 4932  df-br 5167  df-opab 5229  df-mpt 5250  df-id 5593  df-po 5607  df-so 5608  df-xp 5706  df-rel 5707  df-cnv 5708  df-co 5709  df-dm 5710  df-rn 5711  df-res 5712  df-ima 5713  df-iota 6525  df-fun 6575  df-fn 6576  df-f 6577  df-f1 6578  df-fo 6579  df-f1o 6580  df-fv 6581  df-ov 7451  df-oprab 7452  df-mpo 7453  df-er 8763  df-en 9004  df-dom 9005  df-sdom 9006  df-sup 9511  df-pnf 11326  df-mnf 11327  df-ltxr 11329  df-gcd 16541
This theorem is referenced by:  gcdcomd  16560  divgcdnnr  16562  gcdid0  16566  neggcd  16569  gcdabs2  16576  1gcd  16580  6gcd4e2  16585  rprpwr  16606  eucalginv  16631  3lcm2e6woprm  16662  coprmdvds  16700  qredeq  16704  divgcdcoprmex  16713  cncongr1  16714  cncongrprm  16776  fermltl  16831  vfermltl  16848  coprimeprodsq2  16856  pythagtrip  16881  pcgcd  16925  pockthlem  16952  gcdi  17120  gcdmodi  17121  1259lem5  17182  2503lem3  17186  4001lem4  17191  odinv  19603  lgsprme0  27401  lgsdirnn0  27406  lgsquad2lem2  27447  lgsquad3  27449  ex-gcd  30489  gcd32  35711  gcdcomnni  41945  aks6d1c1  42073  aks6d1c4  42081  goldbachthlem2  47420  goldbachth  47421  gcd2odd1  47542  fpprwpprb  47614
  Copyright terms: Public domain W3C validator