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

Theorem coprmdvds 15997
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 16057. (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 11987 . . . . . . . . . 10 (𝑀 ∈ ℤ → 𝑀 ∈ ℂ)
2 zcn 11987 . . . . . . . . . 10 (𝑁 ∈ ℤ → 𝑁 ∈ ℂ)
3 mulcom 10623 . . . . . . . . . 10 ((𝑀 ∈ ℂ ∧ 𝑁 ∈ ℂ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
41, 2, 3syl2an 597 . . . . . . . . 9 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
54breq2d 5078 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · 𝑀)))
6 dvdsmulgcd 15905 . . . . . . . . 9 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
76ancoms 461 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
85, 7bitrd 281 . . . . . . 7 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
983adant1 1126 . . . . . 6 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
109adantr 483 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
11 gcdcom 15862 . . . . . . . . . . 11 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
12113adant3 1128 . . . . . . . . . 10 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
1312eqeq1d 2823 . . . . . . . . 9 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 ↔ (𝑀 gcd 𝐾) = 1))
14 oveq2 7164 . . . . . . . . 9 ((𝑀 gcd 𝐾) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
1513, 14syl6bi 255 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1)))
1615imp 409 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
172mulid1d 10658 . . . . . . . . 9 (𝑁 ∈ ℤ → (𝑁 · 1) = 𝑁)
18173ad2ant3 1131 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 · 1) = 𝑁)
1918adantr 483 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · 1) = 𝑁)
2016, 19eqtrd 2856 . . . . . 6 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = 𝑁)
2120breq2d 5078 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾)) ↔ 𝐾𝑁))
2210, 21bitrd 281 . . . 4 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾𝑁))
2322biimpd 231 . . 3 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁))
2423ex 415 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁)))
2524impcomd 414 1 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 ∥ (𝑀 · 𝑁) ∧ (𝐾 gcd 𝑀) = 1) → 𝐾𝑁))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 208  wa 398  w3a 1083   = wceq 1537  wcel 2114   class class class wbr 5066  (class class class)co 7156  cc 10535  1c1 10538   · cmul 10542  cz 11982  cdvds 15607   gcd cgcd 15843
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1796  ax-4 1810  ax-5 1911  ax-6 1970  ax-7 2015  ax-8 2116  ax-9 2124  ax-10 2145  ax-11 2161  ax-12 2177  ax-ext 2793  ax-sep 5203  ax-nul 5210  ax-pow 5266  ax-pr 5330  ax-un 7461  ax-cnex 10593  ax-resscn 10594  ax-1cn 10595  ax-icn 10596  ax-addcl 10597  ax-addrcl 10598  ax-mulcl 10599  ax-mulrcl 10600  ax-mulcom 10601  ax-addass 10602  ax-mulass 10603  ax-distr 10604  ax-i2m1 10605  ax-1ne0 10606  ax-1rid 10607  ax-rnegex 10608  ax-rrecex 10609  ax-cnre 10610  ax-pre-lttri 10611  ax-pre-lttrn 10612  ax-pre-ltadd 10613  ax-pre-mulgt0 10614  ax-pre-sup 10615
This theorem depends on definitions:  df-bi 209  df-an 399  df-or 844  df-3or 1084  df-3an 1085  df-tru 1540  df-ex 1781  df-nf 1785  df-sb 2070  df-mo 2622  df-eu 2654  df-clab 2800  df-cleq 2814  df-clel 2893  df-nfc 2963  df-ne 3017  df-nel 3124  df-ral 3143  df-rex 3144  df-reu 3145  df-rmo 3146  df-rab 3147  df-v 3496  df-sbc 3773  df-csb 3884  df-dif 3939  df-un 3941  df-in 3943  df-ss 3952  df-pss 3954  df-nul 4292  df-if 4468  df-pw 4541  df-sn 4568  df-pr 4570  df-tp 4572  df-op 4574  df-uni 4839  df-iun 4921  df-br 5067  df-opab 5129  df-mpt 5147  df-tr 5173  df-id 5460  df-eprel 5465  df-po 5474  df-so 5475  df-fr 5514  df-we 5516  df-xp 5561  df-rel 5562  df-cnv 5563  df-co 5564  df-dm 5565  df-rn 5566  df-res 5567  df-ima 5568  df-pred 6148  df-ord 6194  df-on 6195  df-lim 6196  df-suc 6197  df-iota 6314  df-fun 6357  df-fn 6358  df-f 6359  df-f1 6360  df-fo 6361  df-f1o 6362  df-fv 6363  df-riota 7114  df-ov 7159  df-oprab 7160  df-mpo 7161  df-om 7581  df-2nd 7690  df-wrecs 7947  df-recs 8008  df-rdg 8046  df-er 8289  df-en 8510  df-dom 8511  df-sdom 8512  df-sup 8906  df-inf 8907  df-pnf 10677  df-mnf 10678  df-xr 10679  df-ltxr 10680  df-le 10681  df-sub 10872  df-neg 10873  df-div 11298  df-nn 11639  df-2 11701  df-3 11702  df-n0 11899  df-z 11983  df-uz 12245  df-rp 12391  df-fl 13163  df-mod 13239  df-seq 13371  df-exp 13431  df-cj 14458  df-re 14459  df-im 14460  df-sqrt 14594  df-abs 14595  df-dvds 15608  df-gcd 15844
This theorem is referenced by:  coprmdvds2  15998  qredeq  16001  cncongr1  16011  euclemma  16057  eulerthlem2  16119  prmdiveq  16123  prmpwdvds  16240  ablfacrp2  19189  dvdsmulf1o  25771  perfectlem1  25805  lgseisenlem1  25951  lgseisenlem2  25952  lgsquadlem2  25957  lgsquadlem3  25958  2sqlem8  26002  2sqmod  26012  nn0prpwlem  33670  coprmdvdsb  39602  jm2.20nn  39614  perfectALTVlem1  43906
  Copyright terms: Public domain W3C validator