Description: A positive integer divided by the gcd of it and another integer is a positive integer. (Contributed by AV, 10-Jul-2021)
| Ref | Expression | ||
|---|---|---|---|
| Assertion | divgcdnnr | ⊢ ( ( 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ ) → ( 𝐴 / ( 𝐵 gcd 𝐴 ) ) ∈ ℕ ) | 
| Step | Hyp | Ref | Expression | 
|---|---|---|---|
| 1 | nnz | ⊢ ( 𝐴 ∈ ℕ → 𝐴 ∈ ℤ ) | |
| 2 | gcdcom | ⊢ ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ) → ( 𝐴 gcd 𝐵 ) = ( 𝐵 gcd 𝐴 ) ) | |
| 3 | 1 2 | sylan | ⊢ ( ( 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ ) → ( 𝐴 gcd 𝐵 ) = ( 𝐵 gcd 𝐴 ) ) | 
| 4 | 3 | eqcomd | ⊢ ( ( 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ ) → ( 𝐵 gcd 𝐴 ) = ( 𝐴 gcd 𝐵 ) ) | 
| 5 | 4 | oveq2d | ⊢ ( ( 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ ) → ( 𝐴 / ( 𝐵 gcd 𝐴 ) ) = ( 𝐴 / ( 𝐴 gcd 𝐵 ) ) ) | 
| 6 | divgcdnn | ⊢ ( ( 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ ) → ( 𝐴 / ( 𝐴 gcd 𝐵 ) ) ∈ ℕ ) | |
| 7 | 5 6 | eqeltrd | ⊢ ( ( 𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ ) → ( 𝐴 / ( 𝐵 gcd 𝐴 ) ) ∈ ℕ ) |