ILE Home Intuitionistic Logic Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  ILE Home  >  Th. List  >  coprmgcdb GIF version

Theorem coprmgcdb 12102
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 9286 . . . 4 (𝐴 ∈ ℕ → 𝐴 ∈ ℤ)
2 nnz 9286 . . . 4 (𝐵 ∈ ℕ → 𝐵 ∈ ℤ)
3 gcddvds 11978 . . . 4 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
41, 2, 3syl2an 289 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
5 simpr 110 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
6 gcdnncl 11982 . . . . . 6 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 gcd 𝐵) ∈ ℕ)
76adantr 276 . . . . 5 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (𝐴 gcd 𝐵) ∈ ℕ)
8 breq1 4018 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐴 ↔ (𝐴 gcd 𝐵) ∥ 𝐴))
9 breq1 4018 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐵 ↔ (𝐴 gcd 𝐵) ∥ 𝐵))
108, 9anbi12d 473 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → ((𝑖𝐴𝑖𝐵) ↔ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)))
11 eqeq1 2194 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → (𝑖 = 1 ↔ (𝐴 gcd 𝐵) = 1))
1210, 11imbi12d 234 . . . . . 6 (𝑖 = (𝐴 gcd 𝐵) → (((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
1312rspcv 2849 . . . . 5 ((𝐴 gcd 𝐵) ∈ ℕ → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
147, 13syl 14 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
155, 14mpid 42 . . 3 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (𝐴 gcd 𝐵) = 1))
164, 15mpdan 421 . 2 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (𝐴 gcd 𝐵) = 1))
17 simpl 109 . . . . . . . . 9 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) → (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
1817anim1i 340 . . . . . . . 8 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ 𝑖 ∈ ℕ))
1918ancomd 267 . . . . . . 7 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
20 3anass 983 . . . . . . 7 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ↔ (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
2119, 20sylibr 134 . . . . . 6 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
22 nndvdslegcd 11980 . . . . . 6 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
2321, 22syl 14 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
24 breq2 4019 . . . . . . . 8 ((𝐴 gcd 𝐵) = 1 → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
2524adantr 276 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
26 nnge1 8956 . . . . . . . . 9 (𝑖 ∈ ℕ → 1 ≤ 𝑖)
27 nnre 8940 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 𝑖 ∈ ℝ)
28 1red 7986 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 1 ∈ ℝ)
2927, 28letri3d 8087 . . . . . . . . . 10 (𝑖 ∈ ℕ → (𝑖 = 1 ↔ (𝑖 ≤ 1 ∧ 1 ≤ 𝑖)))
3029biimprd 158 . . . . . . . . 9 (𝑖 ∈ ℕ → ((𝑖 ≤ 1 ∧ 1 ≤ 𝑖) → 𝑖 = 1))
3126, 30mpan2d 428 . . . . . . . 8 (𝑖 ∈ ℕ → (𝑖 ≤ 1 → 𝑖 = 1))
3231adantl 277 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ 1 → 𝑖 = 1))
3325, 32sylbid 150 . . . . . 6 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) → 𝑖 = 1))
3433adantll 476 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) → 𝑖 = 1))
3523, 34syld 45 . . . 4 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 = 1))
3635ralrimiva 2560 . . 3 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) → ∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1))
3736ex 115 . 2 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) = 1 → ∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1)))
3816, 37impbid 129 1 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (𝐴 gcd 𝐵) = 1))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 104  wb 105  w3a 979   = wceq 1363  wcel 2158  wral 2465   class class class wbr 4015  (class class class)co 5888  1c1 7826  cle 8007  cn 8933  cz 9267  cdvds 11808   gcd cgcd 11957
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-ia1 106  ax-ia2 107  ax-ia3 108  ax-in1 615  ax-in2 616  ax-io 710  ax-5 1457  ax-7 1458  ax-gen 1459  ax-ie1 1503  ax-ie2 1504  ax-8 1514  ax-10 1515  ax-11 1516  ax-i12 1517  ax-bndl 1519  ax-4 1520  ax-17 1536  ax-i9 1540  ax-ial 1544  ax-i5r 1545  ax-13 2160  ax-14 2161  ax-ext 2169  ax-coll 4130  ax-sep 4133  ax-nul 4141  ax-pow 4186  ax-pr 4221  ax-un 4445  ax-setind 4548  ax-iinf 4599  ax-cnex 7916  ax-resscn 7917  ax-1cn 7918  ax-1re 7919  ax-icn 7920  ax-addcl 7921  ax-addrcl 7922  ax-mulcl 7923  ax-mulrcl 7924  ax-addcom 7925  ax-mulcom 7926  ax-addass 7927  ax-mulass 7928  ax-distr 7929  ax-i2m1 7930  ax-0lt1 7931  ax-1rid 7932  ax-0id 7933  ax-rnegex 7934  ax-precex 7935  ax-cnre 7936  ax-pre-ltirr 7937  ax-pre-ltwlin 7938  ax-pre-lttrn 7939  ax-pre-apti 7940  ax-pre-ltadd 7941  ax-pre-mulgt0 7942  ax-pre-mulext 7943  ax-arch 7944  ax-caucvg 7945
This theorem depends on definitions:  df-bi 117  df-dc 836  df-3or 980  df-3an 981  df-tru 1366  df-fal 1369  df-nf 1471  df-sb 1773  df-eu 2039  df-mo 2040  df-clab 2174  df-cleq 2180  df-clel 2183  df-nfc 2318  df-ne 2358  df-nel 2453  df-ral 2470  df-rex 2471  df-reu 2472  df-rmo 2473  df-rab 2474  df-v 2751  df-sbc 2975  df-csb 3070  df-dif 3143  df-un 3145  df-in 3147  df-ss 3154  df-nul 3435  df-if 3547  df-pw 3589  df-sn 3610  df-pr 3611  df-op 3613  df-uni 3822  df-int 3857  df-iun 3900  df-br 4016  df-opab 4077  df-mpt 4078  df-tr 4114  df-id 4305  df-po 4308  df-iso 4309  df-iord 4378  df-on 4380  df-ilim 4381  df-suc 4383  df-iom 4602  df-xp 4644  df-rel 4645  df-cnv 4646  df-co 4647  df-dm 4648  df-rn 4649  df-res 4650  df-ima 4651  df-iota 5190  df-fun 5230  df-fn 5231  df-f 5232  df-f1 5233  df-fo 5234  df-f1o 5235  df-fv 5236  df-riota 5844  df-ov 5891  df-oprab 5892  df-mpo 5893  df-1st 6155  df-2nd 6156  df-recs 6320  df-frec 6406  df-sup 6997  df-pnf 8008  df-mnf 8009  df-xr 8010  df-ltxr 8011  df-le 8012  df-sub 8144  df-neg 8145  df-reap 8546  df-ap 8553  df-div 8644  df-inn 8934  df-2 8992  df-3 8993  df-4 8994  df-n0 9191  df-z 9268  df-uz 9543  df-q 9634  df-rp 9668  df-fz 10023  df-fzo 10157  df-fl 10284  df-mod 10337  df-seqfrec 10460  df-exp 10534  df-cj 10865  df-re 10866  df-im 10867  df-rsqrt 11021  df-abs 11022  df-dvds 11809  df-gcd 11958
This theorem is referenced by:  coprmdvds1  12105
  Copyright terms: Public domain W3C validator