Metamath Proof Explorer
Theorem nnz
Description: A positive integer is an integer. (Contributed by NM, 9-May-2004)
Reduce dependencies on axioms. (Revised by Steven Nguyen, 29-Nov-2022)
|
|
Ref |
Expression |
|
Assertion |
nnz |
⊢ ( 𝑁 ∈ ℕ → 𝑁 ∈ ℤ ) |
Proof
| Step |
Hyp |
Ref |
Expression |
| 1 |
|
nnre |
⊢ ( 𝑁 ∈ ℕ → 𝑁 ∈ ℝ ) |
| 2 |
|
3mix2 |
⊢ ( 𝑁 ∈ ℕ → ( 𝑁 = 0 ∨ 𝑁 ∈ ℕ ∨ - 𝑁 ∈ ℕ ) ) |
| 3 |
|
elz |
⊢ ( 𝑁 ∈ ℤ ↔ ( 𝑁 ∈ ℝ ∧ ( 𝑁 = 0 ∨ 𝑁 ∈ ℕ ∨ - 𝑁 ∈ ℕ ) ) ) |
| 4 |
1 2 3
|
sylanbrc |
⊢ ( 𝑁 ∈ ℕ → 𝑁 ∈ ℤ ) |