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

Theorem coprmdvds 16087
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 16147. (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 12060 . . . . . . . . . 10 (𝑀 ∈ ℤ → 𝑀 ∈ ℂ)
2 zcn 12060 . . . . . . . . . 10 (𝑁 ∈ ℤ → 𝑁 ∈ ℂ)
3 mulcom 10694 . . . . . . . . . 10 ((𝑀 ∈ ℂ ∧ 𝑁 ∈ ℂ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
41, 2, 3syl2an 599 . . . . . . . . 9 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
54breq2d 5039 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · 𝑀)))
6 dvdsmulgcd 15994 . . . . . . . . 9 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
76ancoms 462 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
85, 7bitrd 282 . . . . . . 7 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
983adant1 1131 . . . . . 6 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
109adantr 484 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
11 gcdcom 15949 . . . . . . . . . . 11 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
12113adant3 1133 . . . . . . . . . 10 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
1312eqeq1d 2740 . . . . . . . . 9 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 ↔ (𝑀 gcd 𝐾) = 1))
14 oveq2 7172 . . . . . . . . 9 ((𝑀 gcd 𝐾) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
1513, 14syl6bi 256 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1)))
1615imp 410 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
172mulid1d 10729 . . . . . . . . 9 (𝑁 ∈ ℤ → (𝑁 · 1) = 𝑁)
18173ad2ant3 1136 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 · 1) = 𝑁)
1918adantr 484 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · 1) = 𝑁)
2016, 19eqtrd 2773 . . . . . 6 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = 𝑁)
2120breq2d 5039 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾)) ↔ 𝐾𝑁))
2210, 21bitrd 282 . . . 4 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾𝑁))
2322biimpd 232 . . 3 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁))
2423ex 416 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁)))
2524impcomd 415 1 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 ∥ (𝑀 · 𝑁) ∧ (𝐾 gcd 𝑀) = 1) → 𝐾𝑁))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 209  wa 399  w3a 1088   = wceq 1542  wcel 2113   class class class wbr 5027  (class class class)co 7164  cc 10606  1c1 10609   · cmul 10613  cz 12055  cdvds 15692   gcd cgcd 15930
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1802  ax-4 1816  ax-5 1916  ax-6 1974  ax-7 2019  ax-8 2115  ax-9 2123  ax-10 2144  ax-11 2161  ax-12 2178  ax-ext 2710  ax-sep 5164  ax-nul 5171  ax-pow 5229  ax-pr 5293  ax-un 7473  ax-cnex 10664  ax-resscn 10665  ax-1cn 10666  ax-icn 10667  ax-addcl 10668  ax-addrcl 10669  ax-mulcl 10670  ax-mulrcl 10671  ax-mulcom 10672  ax-addass 10673  ax-mulass 10674  ax-distr 10675  ax-i2m1 10676  ax-1ne0 10677  ax-1rid 10678  ax-rnegex 10679  ax-rrecex 10680  ax-cnre 10681  ax-pre-lttri 10682  ax-pre-lttrn 10683  ax-pre-ltadd 10684  ax-pre-mulgt0 10685  ax-pre-sup 10686
This theorem depends on definitions:  df-bi 210  df-an 400  df-or 847  df-3or 1089  df-3an 1090  df-tru 1545  df-fal 1555  df-ex 1787  df-nf 1791  df-sb 2074  df-mo 2540  df-eu 2570  df-clab 2717  df-cleq 2730  df-clel 2811  df-nfc 2881  df-ne 2935  df-nel 3039  df-ral 3058  df-rex 3059  df-reu 3060  df-rmo 3061  df-rab 3062  df-v 3399  df-sbc 3680  df-csb 3789  df-dif 3844  df-un 3846  df-in 3848  df-ss 3858  df-pss 3860  df-nul 4210  df-if 4412  df-pw 4487  df-sn 4514  df-pr 4516  df-tp 4518  df-op 4520  df-uni 4794  df-iun 4880  df-br 5028  df-opab 5090  df-mpt 5108  df-tr 5134  df-id 5425  df-eprel 5430  df-po 5438  df-so 5439  df-fr 5478  df-we 5480  df-xp 5525  df-rel 5526  df-cnv 5527  df-co 5528  df-dm 5529  df-rn 5530  df-res 5531  df-ima 5532  df-pred 6123  df-ord 6169  df-on 6170  df-lim 6171  df-suc 6172  df-iota 6291  df-fun 6335  df-fn 6336  df-f 6337  df-f1 6338  df-fo 6339  df-f1o 6340  df-fv 6341  df-riota 7121  df-ov 7167  df-oprab 7168  df-mpo 7169  df-om 7594  df-2nd 7708  df-wrecs 7969  df-recs 8030  df-rdg 8068  df-er 8313  df-en 8549  df-dom 8550  df-sdom 8551  df-sup 8972  df-inf 8973  df-pnf 10748  df-mnf 10749  df-xr 10750  df-ltxr 10751  df-le 10752  df-sub 10943  df-neg 10944  df-div 11369  df-nn 11710  df-2 11772  df-3 11773  df-n0 11970  df-z 12056  df-uz 12318  df-rp 12466  df-fl 13246  df-mod 13322  df-seq 13454  df-exp 13515  df-cj 14541  df-re 14542  df-im 14543  df-sqrt 14677  df-abs 14678  df-dvds 15693  df-gcd 15931
This theorem is referenced by:  coprmdvds2  16088  qredeq  16091  cncongr1  16101  euclemma  16147  eulerthlem2  16212  prmdiveq  16216  prmpwdvds  16333  ablfacrp2  19301  dvdsmulf1o  25923  perfectlem1  25957  lgseisenlem1  26103  lgseisenlem2  26104  lgsquadlem2  26109  lgsquadlem3  26110  2sqlem8  26154  2sqmod  26164  nn0prpwlem  34141  coprmdvdsb  40363  jm2.20nn  40375  perfectALTVlem1  44691
  Copyright terms: Public domain W3C validator