| Intuitionistic Logic Explorer | 
      
      
      < Previous  
      Next >
      
       Nearby theorems  | 
  ||
| Mirrors > Home > ILE Home > Th. List > gcdzeq | GIF version | ||
| Description: A positive integer 𝐴 is equal to its gcd with an integer 𝐵 if and only if 𝐴 divides 𝐵. Generalization of gcdeq 12190. (Contributed by AV, 1-Jul-2020.) | 
| Ref | Expression | 
|---|---|
| gcdzeq | ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) = 𝐴 ↔ 𝐴 ∥ 𝐵)) | 
| Step | Hyp | Ref | Expression | 
|---|---|---|---|
| 1 | nnz 9345 | . . . . 5 ⊢ (𝐴 ∈ ℕ → 𝐴 ∈ ℤ) | |
| 2 | gcddvds 12130 | . . . . 5 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) | |
| 3 | 1, 2 | sylan 283 | . . . 4 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 ∧ (𝐴 gcd 𝐵) ∥ 𝐵)) | 
| 4 | 3 | simprd 114 | . . 3 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ∥ 𝐵) | 
| 5 | breq1 4036 | . . 3 ⊢ ((𝐴 gcd 𝐵) = 𝐴 → ((𝐴 gcd 𝐵) ∥ 𝐵 ↔ 𝐴 ∥ 𝐵)) | |
| 6 | 4, 5 | syl5ibcom 155 | . 2 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) = 𝐴 → 𝐴 ∥ 𝐵)) | 
| 7 | 1 | adantr 276 | . . . . . 6 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → 𝐴 ∈ ℤ) | 
| 8 | iddvds 11969 | . . . . . 6 ⊢ (𝐴 ∈ ℤ → 𝐴 ∥ 𝐴) | |
| 9 | 7, 8 | syl 14 | . . . . 5 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → 𝐴 ∥ 𝐴) | 
| 10 | simpr 110 | . . . . . 6 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → 𝐵 ∈ ℤ) | |
| 11 | nnne0 9018 | . . . . . . . 8 ⊢ (𝐴 ∈ ℕ → 𝐴 ≠ 0) | |
| 12 | simpl 109 | . . . . . . . . 9 ⊢ ((𝐴 = 0 ∧ 𝐵 = 0) → 𝐴 = 0) | |
| 13 | 12 | necon3ai 2416 | . . . . . . . 8 ⊢ (𝐴 ≠ 0 → ¬ (𝐴 = 0 ∧ 𝐵 = 0)) | 
| 14 | 11, 13 | syl 14 | . . . . . . 7 ⊢ (𝐴 ∈ ℕ → ¬ (𝐴 = 0 ∧ 𝐵 = 0)) | 
| 15 | 14 | adantr 276 | . . . . . 6 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ¬ (𝐴 = 0 ∧ 𝐵 = 0)) | 
| 16 | dvdslegcd 12131 | . . . . . 6 ⊢ (((𝐴 ∈ ℤ ∧ 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ ¬ (𝐴 = 0 ∧ 𝐵 = 0)) → ((𝐴 ∥ 𝐴 ∧ 𝐴 ∥ 𝐵) → 𝐴 ≤ (𝐴 gcd 𝐵))) | |
| 17 | 7, 7, 10, 15, 16 | syl31anc 1252 | . . . . 5 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 ∥ 𝐴 ∧ 𝐴 ∥ 𝐵) → 𝐴 ≤ (𝐴 gcd 𝐵))) | 
| 18 | 9, 17 | mpand 429 | . . . 4 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 ∥ 𝐵 → 𝐴 ≤ (𝐴 gcd 𝐵))) | 
| 19 | 3 | simpld 112 | . . . . 5 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ∥ 𝐴) | 
| 20 | gcdcl 12133 | . . . . . . . 8 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ∈ ℕ0) | |
| 21 | 1, 20 | sylan 283 | . . . . . . 7 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ∈ ℕ0) | 
| 22 | 21 | nn0zd 9446 | . . . . . 6 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ∈ ℤ) | 
| 23 | simpl 109 | . . . . . 6 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → 𝐴 ∈ ℕ) | |
| 24 | dvdsle 12009 | . . . . . 6 ⊢ (((𝐴 gcd 𝐵) ∈ ℤ ∧ 𝐴 ∈ ℕ) → ((𝐴 gcd 𝐵) ∥ 𝐴 → (𝐴 gcd 𝐵) ≤ 𝐴)) | |
| 25 | 22, 23, 24 | syl2anc 411 | . . . . 5 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) ∥ 𝐴 → (𝐴 gcd 𝐵) ≤ 𝐴)) | 
| 26 | 19, 25 | mpd 13 | . . . 4 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ≤ 𝐴) | 
| 27 | 18, 26 | jctild 316 | . . 3 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 ∥ 𝐵 → ((𝐴 gcd 𝐵) ≤ 𝐴 ∧ 𝐴 ≤ (𝐴 gcd 𝐵)))) | 
| 28 | 21 | nn0red 9303 | . . . 4 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) ∈ ℝ) | 
| 29 | nnre 8997 | . . . . 5 ⊢ (𝐴 ∈ ℕ → 𝐴 ∈ ℝ) | |
| 30 | 29 | adantr 276 | . . . 4 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → 𝐴 ∈ ℝ) | 
| 31 | 28, 30 | letri3d 8142 | . . 3 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) = 𝐴 ↔ ((𝐴 gcd 𝐵) ≤ 𝐴 ∧ 𝐴 ≤ (𝐴 gcd 𝐵)))) | 
| 32 | 27, 31 | sylibrd 169 | . 2 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴 ∥ 𝐵 → (𝐴 gcd 𝐵) = 𝐴)) | 
| 33 | 6, 32 | impbid 129 | 1 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → ((𝐴 gcd 𝐵) = 𝐴 ↔ 𝐴 ∥ 𝐵)) | 
| Colors of variables: wff set class | 
| Syntax hints: ¬ wn 3 → wi 4 ∧ wa 104 ↔ wb 105 = wceq 1364 ∈ wcel 2167 ≠ wne 2367 class class class wbr 4033 (class class class)co 5922 ℝcr 7878 0cc0 7879 ≤ cle 8062 ℕcn 8990 ℕ0cn0 9249 ℤcz 9326 ∥ cdvds 11952 gcd cgcd 12120 | 
| 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 1461 ax-7 1462 ax-gen 1463 ax-ie1 1507 ax-ie2 1508 ax-8 1518 ax-10 1519 ax-11 1520 ax-i12 1521 ax-bndl 1523 ax-4 1524 ax-17 1540 ax-i9 1544 ax-ial 1548 ax-i5r 1549 ax-13 2169 ax-14 2170 ax-ext 2178 ax-coll 4148 ax-sep 4151 ax-nul 4159 ax-pow 4207 ax-pr 4242 ax-un 4468 ax-setind 4573 ax-iinf 4624 ax-cnex 7970 ax-resscn 7971 ax-1cn 7972 ax-1re 7973 ax-icn 7974 ax-addcl 7975 ax-addrcl 7976 ax-mulcl 7977 ax-mulrcl 7978 ax-addcom 7979 ax-mulcom 7980 ax-addass 7981 ax-mulass 7982 ax-distr 7983 ax-i2m1 7984 ax-0lt1 7985 ax-1rid 7986 ax-0id 7987 ax-rnegex 7988 ax-precex 7989 ax-cnre 7990 ax-pre-ltirr 7991 ax-pre-ltwlin 7992 ax-pre-lttrn 7993 ax-pre-apti 7994 ax-pre-ltadd 7995 ax-pre-mulgt0 7996 ax-pre-mulext 7997 ax-arch 7998 ax-caucvg 7999 | 
| This theorem depends on definitions: df-bi 117 df-dc 836 df-3or 981 df-3an 982 df-tru 1367 df-fal 1370 df-nf 1475 df-sb 1777 df-eu 2048 df-mo 2049 df-clab 2183 df-cleq 2189 df-clel 2192 df-nfc 2328 df-ne 2368 df-nel 2463 df-ral 2480 df-rex 2481 df-reu 2482 df-rmo 2483 df-rab 2484 df-v 2765 df-sbc 2990 df-csb 3085 df-dif 3159 df-un 3161 df-in 3163 df-ss 3170 df-nul 3451 df-if 3562 df-pw 3607 df-sn 3628 df-pr 3629 df-op 3631 df-uni 3840 df-int 3875 df-iun 3918 df-br 4034 df-opab 4095 df-mpt 4096 df-tr 4132 df-id 4328 df-po 4331 df-iso 4332 df-iord 4401 df-on 4403 df-ilim 4404 df-suc 4406 df-iom 4627 df-xp 4669 df-rel 4670 df-cnv 4671 df-co 4672 df-dm 4673 df-rn 4674 df-res 4675 df-ima 4676 df-iota 5219 df-fun 5260 df-fn 5261 df-f 5262 df-f1 5263 df-fo 5264 df-f1o 5265 df-fv 5266 df-riota 5877 df-ov 5925 df-oprab 5926 df-mpo 5927 df-1st 6198 df-2nd 6199 df-recs 6363 df-frec 6449 df-sup 7050 df-pnf 8063 df-mnf 8064 df-xr 8065 df-ltxr 8066 df-le 8067 df-sub 8199 df-neg 8200 df-reap 8602 df-ap 8609 df-div 8700 df-inn 8991 df-2 9049 df-3 9050 df-4 9051 df-n0 9250 df-z 9327 df-uz 9602 df-q 9694 df-rp 9729 df-fz 10084 df-fzo 10218 df-fl 10360 df-mod 10415 df-seqfrec 10540 df-exp 10631 df-cj 11007 df-re 11008 df-im 11009 df-rsqrt 11163 df-abs 11164 df-dvds 11953 df-gcd 12121 | 
| This theorem is referenced by: gcdeq 12190 isevengcd2 12326 | 
| Copyright terms: Public domain | W3C validator |