Metamath Proof Explorer
Description: Nonnegative integer ordering relation. (Contributed by Raph Levien, 10-Dec-2002) (Proof shortened by Mario Carneiro, 16-May-2014)
|
|
Ref |
Expression |
|
Assertion |
nn0ltp1le |
⊢ ( ( 𝑀 ∈ ℕ0 ∧ 𝑁 ∈ ℕ0 ) → ( 𝑀 < 𝑁 ↔ ( 𝑀 + 1 ) ≤ 𝑁 ) ) |
Proof
Step |
Hyp |
Ref |
Expression |
1 |
|
nn0z |
⊢ ( 𝑀 ∈ ℕ0 → 𝑀 ∈ ℤ ) |
2 |
|
nn0z |
⊢ ( 𝑁 ∈ ℕ0 → 𝑁 ∈ ℤ ) |
3 |
|
zltp1le |
⊢ ( ( 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ) → ( 𝑀 < 𝑁 ↔ ( 𝑀 + 1 ) ≤ 𝑁 ) ) |
4 |
1 2 3
|
syl2an |
⊢ ( ( 𝑀 ∈ ℕ0 ∧ 𝑁 ∈ ℕ0 ) → ( 𝑀 < 𝑁 ↔ ( 𝑀 + 1 ) ≤ 𝑁 ) ) |