Description: The addition of two members of a finite set of sequential integers starting at 0 is commutative. (Contributed by Alexander van der Vekens, 22-May-2018) (Revised by Alexander van der Vekens, 9-Jun-2018)
| Ref | Expression | ||
|---|---|---|---|
| Assertion | fz0addcom | ⊢ ( ( 𝐴 ∈ ( 0 ... 𝑁 ) ∧ 𝐵 ∈ ( 0 ... 𝑁 ) ) → ( 𝐴 + 𝐵 ) = ( 𝐵 + 𝐴 ) ) | 
| Step | Hyp | Ref | Expression | 
|---|---|---|---|
| 1 | elfznn0 | ⊢ ( 𝐴 ∈ ( 0 ... 𝑁 ) → 𝐴 ∈ ℕ0 ) | |
| 2 | 1 | nn0cnd | ⊢ ( 𝐴 ∈ ( 0 ... 𝑁 ) → 𝐴 ∈ ℂ ) | 
| 3 | elfznn0 | ⊢ ( 𝐵 ∈ ( 0 ... 𝑁 ) → 𝐵 ∈ ℕ0 ) | |
| 4 | 3 | nn0cnd | ⊢ ( 𝐵 ∈ ( 0 ... 𝑁 ) → 𝐵 ∈ ℂ ) | 
| 5 | addcom | ⊢ ( ( 𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ) → ( 𝐴 + 𝐵 ) = ( 𝐵 + 𝐴 ) ) | |
| 6 | 2 4 5 | syl2an | ⊢ ( ( 𝐴 ∈ ( 0 ... 𝑁 ) ∧ 𝐵 ∈ ( 0 ... 𝑁 ) ) → ( 𝐴 + 𝐵 ) = ( 𝐵 + 𝐴 ) ) |