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

Theorem coprmdvds 16677
Description: Euclid's Lemma (see ProofWiki "Euclid's Lemma", 10-Jul-2021, https://proofwiki.org/wiki/Euclid's_Lemma): If an integer divides the product of two integers and is coprime to one of them, then it divides the other. See also theorem 1.5 in [ApostolNT] p. 16. Generalization of euclemma 16737. (Contributed by Paul Chapman, 22-Jun-2011.) (Proof shortened by AV, 10-Jul-2021.)
Assertion
Ref Expression
coprmdvds ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 ∥ (𝑀 · 𝑁) ∧ (𝐾 gcd 𝑀) = 1) → 𝐾𝑁))

Proof of Theorem coprmdvds
StepHypRef Expression
1 zcn 12598 . . . . . . . . . 10 (𝑀 ∈ ℤ → 𝑀 ∈ ℂ)
2 zcn 12598 . . . . . . . . . 10 (𝑁 ∈ ℤ → 𝑁 ∈ ℂ)
3 mulcom 11220 . . . . . . . . . 10 ((𝑀 ∈ ℂ ∧ 𝑁 ∈ ℂ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
41, 2, 3syl2an 596 . . . . . . . . 9 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
54breq2d 5136 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · 𝑀)))
6 dvdsmulgcd 16580 . . . . . . . . 9 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
76ancoms 458 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
85, 7bitrd 279 . . . . . . 7 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
983adant1 1130 . . . . . 6 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
109adantr 480 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
11 gcdcom 16537 . . . . . . . . . . 11 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
12113adant3 1132 . . . . . . . . . 10 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
1312eqeq1d 2738 . . . . . . . . 9 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 ↔ (𝑀 gcd 𝐾) = 1))
14 oveq2 7418 . . . . . . . . 9 ((𝑀 gcd 𝐾) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
1513, 14biimtrdi 253 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1)))
1615imp 406 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
172mulridd 11257 . . . . . . . . 9 (𝑁 ∈ ℤ → (𝑁 · 1) = 𝑁)
18173ad2ant3 1135 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 · 1) = 𝑁)
1918adantr 480 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · 1) = 𝑁)
2016, 19eqtrd 2771 . . . . . 6 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = 𝑁)
2120breq2d 5136 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾)) ↔ 𝐾𝑁))
2210, 21bitrd 279 . . . 4 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾𝑁))
2322biimpd 229 . . 3 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁))
2423ex 412 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁)))
2524impcomd 411 1 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 ∥ (𝑀 · 𝑁) ∧ (𝐾 gcd 𝑀) = 1) → 𝐾𝑁))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 206  wa 395  w3a 1086   = wceq 1540  wcel 2109   class class class wbr 5124  (class class class)co 7410  cc 11132  1c1 11135   · cmul 11139  cz 12593  cdvds 16277   gcd cgcd 16518
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 2008  ax-8 2111  ax-9 2119  ax-10 2142  ax-11 2158  ax-12 2178  ax-ext 2708  ax-sep 5271  ax-nul 5281  ax-pow 5340  ax-pr 5407  ax-un 7734  ax-cnex 11190  ax-resscn 11191  ax-1cn 11192  ax-icn 11193  ax-addcl 11194  ax-addrcl 11195  ax-mulcl 11196  ax-mulrcl 11197  ax-mulcom 11198  ax-addass 11199  ax-mulass 11200  ax-distr 11201  ax-i2m1 11202  ax-1ne0 11203  ax-1rid 11204  ax-rnegex 11205  ax-rrecex 11206  ax-cnre 11207  ax-pre-lttri 11208  ax-pre-lttrn 11209  ax-pre-ltadd 11210  ax-pre-mulgt0 11211  ax-pre-sup 11212
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 2066  df-mo 2540  df-eu 2569  df-clab 2715  df-cleq 2728  df-clel 2810  df-nfc 2886  df-ne 2934  df-nel 3038  df-ral 3053  df-rex 3062  df-rmo 3364  df-reu 3365  df-rab 3421  df-v 3466  df-sbc 3771  df-csb 3880  df-dif 3934  df-un 3936  df-in 3938  df-ss 3948  df-pss 3951  df-nul 4314  df-if 4506  df-pw 4582  df-sn 4607  df-pr 4609  df-op 4613  df-uni 4889  df-iun 4974  df-br 5125  df-opab 5187  df-mpt 5207  df-tr 5235  df-id 5553  df-eprel 5558  df-po 5566  df-so 5567  df-fr 5611  df-we 5613  df-xp 5665  df-rel 5666  df-cnv 5667  df-co 5668  df-dm 5669  df-rn 5670  df-res 5671  df-ima 5672  df-pred 6295  df-ord 6360  df-on 6361  df-lim 6362  df-suc 6363  df-iota 6489  df-fun 6538  df-fn 6539  df-f 6540  df-f1 6541  df-fo 6542  df-f1o 6543  df-fv 6544  df-riota 7367  df-ov 7413  df-oprab 7414  df-mpo 7415  df-om 7867  df-2nd 7994  df-frecs 8285  df-wrecs 8316  df-recs 8390  df-rdg 8429  df-er 8724  df-en 8965  df-dom 8966  df-sdom 8967  df-sup 9459  df-inf 9460  df-pnf 11276  df-mnf 11277  df-xr 11278  df-ltxr 11279  df-le 11280  df-sub 11473  df-neg 11474  df-div 11900  df-nn 12246  df-2 12308  df-3 12309  df-n0 12507  df-z 12594  df-uz 12858  df-rp 13014  df-fl 13814  df-mod 13892  df-seq 14025  df-exp 14085  df-cj 15123  df-re 15124  df-im 15125  df-sqrt 15259  df-abs 15260  df-dvds 16278  df-gcd 16519
This theorem is referenced by:  coprmdvds2  16678  qredeq  16681  cncongr1  16691  euclemma  16737  eulerthlem2  16806  prmdiveq  16810  prmpwdvds  16929  ablfacrp2  20055  mpodvdsmulf1o  27161  dvdsmulf1o  27163  perfectlem1  27197  lgseisenlem1  27343  lgseisenlem2  27344  lgsquadlem2  27349  lgsquadlem3  27350  2sqlem8  27394  2sqmod  27404  nn0prpwlem  36345  hashscontpow1  42139  coprmdvdsb  42976  jm2.20nn  42988  perfectALTVlem1  47702
  Copyright terms: Public domain W3C validator