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

Theorem coprmgcdb 12088
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 9272 . . . 4 (𝐴 ∈ ℕ → 𝐴 ∈ ℤ)
2 nnz 9272 . . . 4 (𝐵 ∈ ℕ → 𝐵 ∈ ℤ)
3 gcddvds 11964 . . . 4 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
41, 2, 3syl2an 289 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
5 simpr 110 . . . 4 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵))
6 gcdnncl 11968 . . . . . 6 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 gcd 𝐵) ∈ ℕ)
76adantr 276 . . . . 5 (((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) → (𝐴 gcd 𝐵) ∈ ℕ)
8 breq1 4007 . . . . . . . 8 (𝑖 = (𝐴 gcd 𝐵) → (𝑖𝐴 ↔ (𝐴 gcd 𝐵) ∥ 𝐴))
9 breq1 4007 . . . . . . . 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 2838 . . . . 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 11966 . . . . . 6 ((𝑖 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
2321, 22syl 14 . . . . 5 ((((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) ∧ (𝐴 gcd 𝐵) = 1) ∧ 𝑖 ∈ ℕ) → ((𝑖𝐴𝑖𝐵) → 𝑖 ≤ (𝐴 gcd 𝐵)))
24 breq2 4008 . . . . . . . 8 ((𝐴 gcd 𝐵) = 1 → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
2524adantr 276 . . . . . . 7 (((𝐴 gcd 𝐵) = 1 ∧ 𝑖 ∈ ℕ) → (𝑖 ≤ (𝐴 gcd 𝐵) ↔ 𝑖 ≤ 1))
26 nnge1 8942 . . . . . . . . 9 (𝑖 ∈ ℕ → 1 ≤ 𝑖)
27 nnre 8926 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 𝑖 ∈ ℝ)
28 1red 7972 . . . . . . . . . . 11 (𝑖 ∈ ℕ → 1 ∈ ℝ)
2927, 28letri3d 8073 . . . . . . . . . 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 4004  (class class class)co 5875  1c1 7812  cle 7993  cn 8919  cz 9253  cdvds 11794   gcd cgcd 11943
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 4119  ax-sep 4122  ax-nul 4130  ax-pow 4175  ax-pr 4210  ax-un 4434  ax-setind 4537  ax-iinf 4588  ax-cnex 7902  ax-resscn 7903  ax-1cn 7904  ax-1re 7905  ax-icn 7906  ax-addcl 7907  ax-addrcl 7908  ax-mulcl 7909  ax-mulrcl 7910  ax-addcom 7911  ax-mulcom 7912  ax-addass 7913  ax-mulass 7914  ax-distr 7915  ax-i2m1 7916  ax-0lt1 7917  ax-1rid 7918  ax-0id 7919  ax-rnegex 7920  ax-precex 7921  ax-cnre 7922  ax-pre-ltirr 7923  ax-pre-ltwlin 7924  ax-pre-lttrn 7925  ax-pre-apti 7926  ax-pre-ltadd 7927  ax-pre-mulgt0 7928  ax-pre-mulext 7929  ax-arch 7930  ax-caucvg 7931
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 2740  df-sbc 2964  df-csb 3059  df-dif 3132  df-un 3134  df-in 3136  df-ss 3143  df-nul 3424  df-if 3536  df-pw 3578  df-sn 3599  df-pr 3600  df-op 3602  df-uni 3811  df-int 3846  df-iun 3889  df-br 4005  df-opab 4066  df-mpt 4067  df-tr 4103  df-id 4294  df-po 4297  df-iso 4298  df-iord 4367  df-on 4369  df-ilim 4370  df-suc 4372  df-iom 4591  df-xp 4633  df-rel 4634  df-cnv 4635  df-co 4636  df-dm 4637  df-rn 4638  df-res 4639  df-ima 4640  df-iota 5179  df-fun 5219  df-fn 5220  df-f 5221  df-f1 5222  df-fo 5223  df-f1o 5224  df-fv 5225  df-riota 5831  df-ov 5878  df-oprab 5879  df-mpo 5880  df-1st 6141  df-2nd 6142  df-recs 6306  df-frec 6392  df-sup 6983  df-pnf 7994  df-mnf 7995  df-xr 7996  df-ltxr 7997  df-le 7998  df-sub 8130  df-neg 8131  df-reap 8532  df-ap 8539  df-div 8630  df-inn 8920  df-2 8978  df-3 8979  df-4 8980  df-n0 9177  df-z 9254  df-uz 9529  df-q 9620  df-rp 9654  df-fz 10009  df-fzo 10143  df-fl 10270  df-mod 10323  df-seqfrec 10446  df-exp 10520  df-cj 10851  df-re 10852  df-im 10853  df-rsqrt 11007  df-abs 11008  df-dvds 11795  df-gcd 11944
This theorem is referenced by:  coprmdvds1  12091
  Copyright terms: Public domain W3C validator