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

Theorem coprmgcdb 12071
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 9261 . . . 4 (𝐴 ∈ ℕ → 𝐴 ∈ ℤ)
2 nnz 9261 . . . 4 (𝐵 ∈ ℕ → 𝐵 ∈ ℤ)
3 gcddvds 11947 . . . 4 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
41, 2, 3syl2an 289 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
5 simpr 110 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
6 gcdnncl 11951 . . . . . 6 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 gcd 𝐵) ∈ ℕ)
76adantr 276 . . . . 5 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (𝐴 gcd 𝐵) ∈ ℕ)
8 breq1 4003 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐴 ↔ (𝐴 gcd 𝐵) ∥ 𝐴))
9 breq1 4003 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐵 ↔ (𝐴 gcd 𝐵) ∥ 𝐵))
108, 9anbi12d 473 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → ((𝑖𝐴𝑖𝐵) ↔ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)))
11 eqeq1 2184 . . . . . . 7 (𝑖 = (𝐴 gcd 𝐵) → (𝑖 = 1 ↔ (𝐴 gcd 𝐵) = 1))
1210, 11imbi12d 234 . . . . . 6 (𝑖 = (𝐴 gcd 𝐵) → (((𝑖𝐴𝑖𝐵) → 𝑖 = 1) ↔ (((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵) → (𝐴 gcd 𝐵) = 1)))
1312rspcv 2837 . . . . 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 982 . . . . . . 7 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ↔ (𝑖 ∈ ℕ ∧ (𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ)))
2119, 20sylibr 134 . . . . . 6 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → (𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ))
22 nndvdslegcd 11949 . . . . . 6 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
2321, 22syl 14 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
24 breq2 4004 . . . . . . . 8 ((𝐴 gcd 𝐵) = 1 → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
2524adantr 276 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
26 nnge1 8931 . . . . . . . . 9 (𝑖 ∈ ℕ → 1 ≤ 𝑖)
27 nnre 8915 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 𝑖 ∈ ℝ)
28 1red 7963 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 1 ∈ ℝ)
2927, 28letri3d 8063 . . . . . . . . . 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 2550 . . 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 978   = wceq 1353  wcel 2148  wral 2455   class class class wbr 4000  (class class class)co 5869  1c1 7803  cle 7983  cn 8908  cz 9242  cdvds 11778   gcd cgcd 11926
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 614  ax-in2 615  ax-io 709  ax-5 1447  ax-7 1448  ax-gen 1449  ax-ie1 1493  ax-ie2 1494  ax-8 1504  ax-10 1505  ax-11 1506  ax-i12 1507  ax-bndl 1509  ax-4 1510  ax-17 1526  ax-i9 1530  ax-ial 1534  ax-i5r 1535  ax-13 2150  ax-14 2151  ax-ext 2159  ax-coll 4115  ax-sep 4118  ax-nul 4126  ax-pow 4171  ax-pr 4206  ax-un 4430  ax-setind 4533  ax-iinf 4584  ax-cnex 7893  ax-resscn 7894  ax-1cn 7895  ax-1re 7896  ax-icn 7897  ax-addcl 7898  ax-addrcl 7899  ax-mulcl 7900  ax-mulrcl 7901  ax-addcom 7902  ax-mulcom 7903  ax-addass 7904  ax-mulass 7905  ax-distr 7906  ax-i2m1 7907  ax-0lt1 7908  ax-1rid 7909  ax-0id 7910  ax-rnegex 7911  ax-precex 7912  ax-cnre 7913  ax-pre-ltirr 7914  ax-pre-ltwlin 7915  ax-pre-lttrn 7916  ax-pre-apti 7917  ax-pre-ltadd 7918  ax-pre-mulgt0 7919  ax-pre-mulext 7920  ax-arch 7921  ax-caucvg 7922
This theorem depends on definitions:  df-bi 117  df-dc 835  df-3or 979  df-3an 980  df-tru 1356  df-fal 1359  df-nf 1461  df-sb 1763  df-eu 2029  df-mo 2030  df-clab 2164  df-cleq 2170  df-clel 2173  df-nfc 2308  df-ne 2348  df-nel 2443  df-ral 2460  df-rex 2461  df-reu 2462  df-rmo 2463  df-rab 2464  df-v 2739  df-sbc 2963  df-csb 3058  df-dif 3131  df-un 3133  df-in 3135  df-ss 3142  df-nul 3423  df-if 3535  df-pw 3576  df-sn 3597  df-pr 3598  df-op 3600  df-uni 3808  df-int 3843  df-iun 3886  df-br 4001  df-opab 4062  df-mpt 4063  df-tr 4099  df-id 4290  df-po 4293  df-iso 4294  df-iord 4363  df-on 4365  df-ilim 4366  df-suc 4368  df-iom 4587  df-xp 4629  df-rel 4630  df-cnv 4631  df-co 4632  df-dm 4633  df-rn 4634  df-res 4635  df-ima 4636  df-iota 5174  df-fun 5214  df-fn 5215  df-f 5216  df-f1 5217  df-fo 5218  df-f1o 5219  df-fv 5220  df-riota 5825  df-ov 5872  df-oprab 5873  df-mpo 5874  df-1st 6135  df-2nd 6136  df-recs 6300  df-frec 6386  df-sup 6977  df-pnf 7984  df-mnf 7985  df-xr 7986  df-ltxr 7987  df-le 7988  df-sub 8120  df-neg 8121  df-reap 8522  df-ap 8529  df-div 8619  df-inn 8909  df-2 8967  df-3 8968  df-4 8969  df-n0 9166  df-z 9243  df-uz 9518  df-q 9609  df-rp 9641  df-fz 9996  df-fzo 10129  df-fl 10256  df-mod 10309  df-seqfrec 10432  df-exp 10506  df-cj 10835  df-re 10836  df-im 10837  df-rsqrt 10991  df-abs 10992  df-dvds 11779  df-gcd 11927
This theorem is referenced by:  coprmdvds1  12074
  Copyright terms: Public domain W3C validator