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

Theorem coprmgcdb 15768
Description: Two positive integers are coprime, i.e. the only positive integer that divides both of them is 1, iff their greatest common divisor is 1. (Contributed by AV, 9-Aug-2020.)
Assertion
Ref Expression
coprmgcdb ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (𝐴 gcd 𝐵) = 1))
Distinct variable groups:   𝐴,𝑖   𝐵,𝑖

Proof of Theorem coprmgcdb
StepHypRef Expression
1 nnz 11751 . . . 4 (𝐴 ∈ ℕ → 𝐴 ∈ ℤ)
2 nnz 11751 . . . 4 (𝐵 ∈ ℕ → 𝐵 ∈ ℤ)
3 gcddvds 15631 . . . 4 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
41, 2, 3syl2an 589 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
5 simpr 479 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
6 gcdnncl 15635 . . . . . 6 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 gcd 𝐵) ∈ ℕ)
76adantr 474 . . . . 5 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (𝐴 gcd 𝐵) ∈ ℕ)
8 breq1 4889 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐴 ↔ (𝐴 gcd 𝐵) ∥ 𝐴))
9 breq1 4889 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐵 ↔ (𝐴 gcd 𝐵) ∥ 𝐵))
108, 9anbi12d 624 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → ((𝑖𝐴𝑖𝐵) ↔ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)))
11 eqeq1 2782 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → (𝑖 = 1 ↔ (𝐴 gcd 𝐵) = 1))
1210, 11imbi12d 336 . . . . . 6 (𝑖 = (𝐴 gcd 𝐵) → (((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
1312rspcv 3507 . . . . 5 ((𝐴 gcd 𝐵) ∈ ℕ → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
147, 13syl 17 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
155, 14mpid 44 . . 3 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (𝐴 gcd 𝐵) = 1))
164, 15mpdan 677 . 2 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (𝐴 gcd 𝐵) = 1))
17 simpl 476 . . . . . . . 8 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) → (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
1817anim1ci 609 . . . . . . 7 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
19 3anass 1079 . . . . . . 7 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ↔ (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
2018, 19sylibr 226 . . . . . 6 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
21 nndvdslegcd 15633 . . . . . 6 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
2220, 21syl 17 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
23 breq2 4890 . . . . . . . 8 ((𝐴 gcd 𝐵) = 1 → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
2423adantr 474 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
25 nnge1 11404 . . . . . . . . 9 (𝑖 ∈ ℕ → 1 ≤ 𝑖)
26 nnre 11382 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 𝑖 ∈ ℝ)
27 1red 10377 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 1 ∈ ℝ)
2826, 27letri3d 10518 . . . . . . . . . 10 (𝑖 ∈ ℕ → (𝑖 = 1 ↔ (𝑖 ≤ 1 ∧ 1 ≤ 𝑖)))
2928biimprd 240 . . . . . . . . 9 (𝑖 ∈ ℕ → ((𝑖 ≤ 1 ∧ 1 ≤ 𝑖) → 𝑖 = 1))
3025, 29mpan2d 684 . . . . . . . 8 (𝑖 ∈ ℕ → (𝑖 ≤ 1 → 𝑖 = 1))
3130adantl 475 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ 1 → 𝑖 = 1))
3224, 31sylbid 232 . . . . . 6 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) → 𝑖 = 1))
3332adantll 704 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) → 𝑖 = 1))
3422, 33syld 47 . . . 4 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 = 1))
3534ralrimiva 3148 . . 3 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) → ∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1))
3635ex 403 . 2 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) = 1 → ∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1)))
3716, 36impbid 204 1 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (𝐴 gcd 𝐵) = 1))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 198  wa 386  w3a 1071   = wceq 1601  wcel 2107  wral 3090   class class class wbr 4886  (class class class)co 6922  1c1 10273  cle 10412  cn 11374  cz 11728  cdvds 15387   gcd cgcd 15622
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1839  ax-4 1853  ax-5 1953  ax-6 2021  ax-7 2055  ax-8 2109  ax-9 2116  ax-10 2135  ax-11 2150  ax-12 2163  ax-13 2334  ax-ext 2754  ax-sep 5017  ax-nul 5025  ax-pow 5077  ax-pr 5138  ax-un 7226  ax-cnex 10328  ax-resscn 10329  ax-1cn 10330  ax-icn 10331  ax-addcl 10332  ax-addrcl 10333  ax-mulcl 10334  ax-mulrcl 10335  ax-mulcom 10336  ax-addass 10337  ax-mulass 10338  ax-distr 10339  ax-i2m1 10340  ax-1ne0 10341  ax-1rid 10342  ax-rnegex 10343  ax-rrecex 10344  ax-cnre 10345  ax-pre-lttri 10346  ax-pre-lttrn 10347  ax-pre-ltadd 10348  ax-pre-mulgt0 10349  ax-pre-sup 10350
This theorem depends on definitions:  df-bi 199  df-an 387  df-or 837  df-3or 1072  df-3an 1073  df-tru 1605  df-ex 1824  df-nf 1828  df-sb 2012  df-mo 2551  df-eu 2587  df-clab 2764  df-cleq 2770  df-clel 2774  df-nfc 2921  df-ne 2970  df-nel 3076  df-ral 3095  df-rex 3096  df-reu 3097  df-rmo 3098  df-rab 3099  df-v 3400  df-sbc 3653  df-csb 3752  df-dif 3795  df-un 3797  df-in 3799  df-ss 3806  df-pss 3808  df-nul 4142  df-if 4308  df-pw 4381  df-sn 4399  df-pr 4401  df-tp 4403  df-op 4405  df-uni 4672  df-iun 4755  df-br 4887  df-opab 4949  df-mpt 4966  df-tr 4988  df-id 5261  df-eprel 5266  df-po 5274  df-so 5275  df-fr 5314  df-we 5316  df-xp 5361  df-rel 5362  df-cnv 5363  df-co 5364  df-dm 5365  df-rn 5366  df-res 5367  df-ima 5368  df-pred 5933  df-ord 5979  df-on 5980  df-lim 5981  df-suc 5982  df-iota 6099  df-fun 6137  df-fn 6138  df-f 6139  df-f1 6140  df-fo 6141  df-f1o 6142  df-fv 6143  df-riota 6883  df-ov 6925  df-oprab 6926  df-mpt2 6927  df-om 7344  df-2nd 7446  df-wrecs 7689  df-recs 7751  df-rdg 7789  df-er 8026  df-en 8242  df-dom 8243  df-sdom 8244  df-sup 8636  df-inf 8637  df-pnf 10413  df-mnf 10414  df-xr 10415  df-ltxr 10416  df-le 10417  df-sub 10608  df-neg 10609  df-div 11033  df-nn 11375  df-2 11438  df-3 11439  df-n0 11643  df-z 11729  df-uz 11993  df-rp 12138  df-seq 13120  df-exp 13179  df-cj 14246  df-re 14247  df-im 14248  df-sqrt 14382  df-abs 14383  df-dvds 15388  df-gcd 15623
This theorem is referenced by:  ncoprmgcdne1b  15769  coprmdvds1  15771
  Copyright terms: Public domain W3C validator