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

Theorem coprmgcdb 11615
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 8977 . . . 4 (𝐴 ∈ ℕ → 𝐴 ∈ ℤ)
2 nnz 8977 . . . 4 (𝐵 ∈ ℕ → 𝐵 ∈ ℤ)
3 gcddvds 11500 . . . 4 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
41, 2, 3syl2an 285 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
5 simpr 109 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
6 gcdnncl 11504 . . . . . 6 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 gcd 𝐵) ∈ ℕ)
76adantr 272 . . . . 5 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (𝐴 gcd 𝐵) ∈ ℕ)
8 breq1 3898 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐴 ↔ (𝐴 gcd 𝐵) ∥ 𝐴))
9 breq1 3898 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐵 ↔ (𝐴 gcd 𝐵) ∥ 𝐵))
108, 9anbi12d 462 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → ((𝑖𝐴𝑖𝐵) ↔ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)))
11 eqeq1 2121 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → (𝑖 = 1 ↔ (𝐴 gcd 𝐵) = 1))
1210, 11imbi12d 233 . . . . . 6 (𝑖 = (𝐴 gcd 𝐵) → (((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
1312rspcv 2756 . . . . 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 415 . 2 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) → (𝐴 gcd 𝐵) = 1))
17 simpl 108 . . . . . . . . 9 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) → (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
1817anim1i 336 . . . . . . . 8 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ 𝑖 ∈ ℕ))
1918ancomd 265 . . . . . . 7 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
20 3anass 949 . . . . . . 7 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ↔ (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
2119, 20sylibr 133 . . . . . 6 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
22 nndvdslegcd 11502 . . . . . 6 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
2321, 22syl 14 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
24 breq2 3899 . . . . . . . 8 ((𝐴 gcd 𝐵) = 1 → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
2524adantr 272 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
26 nnge1 8653 . . . . . . . . 9 (𝑖 ∈ ℕ → 1 ≤ 𝑖)
27 nnre 8637 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 𝑖 ∈ ℝ)
28 1red 7705 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 1 ∈ ℝ)
2927, 28letri3d 7802 . . . . . . . . . 10 (𝑖 ∈ ℕ → (𝑖 = 1 ↔ (𝑖 ≤ 1 ∧ 1 ≤ 𝑖)))
3029biimprd 157 . . . . . . . . 9 (𝑖 ∈ ℕ → ((𝑖 ≤ 1 ∧ 1 ≤ 𝑖) → 𝑖 = 1))
3126, 30mpan2d 422 . . . . . . . 8 (𝑖 ∈ ℕ → (𝑖 ≤ 1 → 𝑖 = 1))
3231adantl 273 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ 1 → 𝑖 = 1))
3325, 32sylbid 149 . . . . . 6 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) → 𝑖 = 1))
3433adantll 465 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) → 𝑖 = 1))
3523, 34syld 45 . . . 4 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 = 1))
3635ralrimiva 2479 . . 3 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) → ∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1))
3736ex 114 . 2 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) = 1 → ∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1)))
3816, 37impbid 128 1 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (∀𝑖 ∈ ℕ ((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (𝐴 gcd 𝐵) = 1))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  wb 104  w3a 945   = wceq 1314  wcel 1463  wral 2390   class class class wbr 3895  (class class class)co 5728  1c1 7548  cle 7725  cn 8630  cz 8958  cdvds 11341   gcd cgcd 11483
This theorem was proved from axioms:  ax-1 5  ax-2 6  ax-mp 7  ax-ia1 105  ax-ia2 106  ax-ia3 107  ax-in1 586  ax-in2 587  ax-io 681  ax-5 1406  ax-7 1407  ax-gen 1408  ax-ie1 1452  ax-ie2 1453  ax-8 1465  ax-10 1466  ax-11 1467  ax-i12 1468  ax-bndl 1469  ax-4 1470  ax-13 1474  ax-14 1475  ax-17 1489  ax-i9 1493  ax-ial 1497  ax-i5r 1498  ax-ext 2097  ax-coll 4003  ax-sep 4006  ax-nul 4014  ax-pow 4058  ax-pr 4091  ax-un 4315  ax-setind 4412  ax-iinf 4462  ax-cnex 7636  ax-resscn 7637  ax-1cn 7638  ax-1re 7639  ax-icn 7640  ax-addcl 7641  ax-addrcl 7642  ax-mulcl 7643  ax-mulrcl 7644  ax-addcom 7645  ax-mulcom 7646  ax-addass 7647  ax-mulass 7648  ax-distr 7649  ax-i2m1 7650  ax-0lt1 7651  ax-1rid 7652  ax-0id 7653  ax-rnegex 7654  ax-precex 7655  ax-cnre 7656  ax-pre-ltirr 7657  ax-pre-ltwlin 7658  ax-pre-lttrn 7659  ax-pre-apti 7660  ax-pre-ltadd 7661  ax-pre-mulgt0 7662  ax-pre-mulext 7663  ax-arch 7664  ax-caucvg 7665
This theorem depends on definitions:  df-bi 116  df-dc 803  df-3or 946  df-3an 947  df-tru 1317  df-fal 1320  df-nf 1420  df-sb 1719  df-eu 1978  df-mo 1979  df-clab 2102  df-cleq 2108  df-clel 2111  df-nfc 2244  df-ne 2283  df-nel 2378  df-ral 2395  df-rex 2396  df-reu 2397  df-rmo 2398  df-rab 2399  df-v 2659  df-sbc 2879  df-csb 2972  df-dif 3039  df-un 3041  df-in 3043  df-ss 3050  df-nul 3330  df-if 3441  df-pw 3478  df-sn 3499  df-pr 3500  df-op 3502  df-uni 3703  df-int 3738  df-iun 3781  df-br 3896  df-opab 3950  df-mpt 3951  df-tr 3987  df-id 4175  df-po 4178  df-iso 4179  df-iord 4248  df-on 4250  df-ilim 4251  df-suc 4253  df-iom 4465  df-xp 4505  df-rel 4506  df-cnv 4507  df-co 4508  df-dm 4509  df-rn 4510  df-res 4511  df-ima 4512  df-iota 5046  df-fun 5083  df-fn 5084  df-f 5085  df-f1 5086  df-fo 5087  df-f1o 5088  df-fv 5089  df-riota 5684  df-ov 5731  df-oprab 5732  df-mpo 5733  df-1st 5992  df-2nd 5993  df-recs 6156  df-frec 6242  df-sup 6823  df-pnf 7726  df-mnf 7727  df-xr 7728  df-ltxr 7729  df-le 7730  df-sub 7858  df-neg 7859  df-reap 8255  df-ap 8262  df-div 8346  df-inn 8631  df-2 8689  df-3 8690  df-4 8691  df-n0 8882  df-z 8959  df-uz 9229  df-q 9314  df-rp 9344  df-fz 9684  df-fzo 9813  df-fl 9936  df-mod 9989  df-seqfrec 10112  df-exp 10186  df-cj 10507  df-re 10508  df-im 10509  df-rsqrt 10662  df-abs 10663  df-dvds 11342  df-gcd 11484
This theorem is referenced by:  coprmdvds1  11618
  Copyright terms: Public domain W3C validator