Metamath Proof Explorer


Theorem isodd2

Description: The predicate "is an odd number". An odd number is an integer which is not divisible by 2, i.e. the result of dividing the odd number decreased by 1 and then divided by 2 is still an integer. (Contributed by AV, 15-Jun-2020)

Ref Expression
Assertion isodd2 ( 𝑍 ∈ Odd ↔ ( 𝑍 ∈ ℤ ∧ ( ( 𝑍 − 1 ) / 2 ) ∈ ℤ ) )

Proof

Step Hyp Ref Expression
1 isodd ( 𝑍 ∈ Odd ↔ ( 𝑍 ∈ ℤ ∧ ( ( 𝑍 + 1 ) / 2 ) ∈ ℤ ) )
2 zob ( 𝑍 ∈ ℤ → ( ( ( 𝑍 + 1 ) / 2 ) ∈ ℤ ↔ ( ( 𝑍 − 1 ) / 2 ) ∈ ℤ ) )
3 2 pm5.32i ( ( 𝑍 ∈ ℤ ∧ ( ( 𝑍 + 1 ) / 2 ) ∈ ℤ ) ↔ ( 𝑍 ∈ ℤ ∧ ( ( 𝑍 − 1 ) / 2 ) ∈ ℤ ) )
4 1 3 bitri ( 𝑍 ∈ Odd ↔ ( 𝑍 ∈ ℤ ∧ ( ( 𝑍 − 1 ) / 2 ) ∈ ℤ ) )