Bitwise operations on natural numbers
In the first half of this file, we provide theorems for reasoning about natural numbers from their
bitwise properties. In the second half of this file, we show properties of the bitwise operations
lor
, land
and lxor
, which are defined in core.
Main results
eq_of_test_bit_eq
: two natural numbers are equal if they have equal bits at every position.exists_most_significant_bit
: ifn ≠ 0
, then there is some positioni
that contains the most significant1
-bit ofn
.lt_of_test_bit
: ifn
andm
are numbers andi
is a position such that thei
-th bit of ofn
is zero, thei
-th bit ofm
is one, and all more significant bits are equal, thenn < m
.
Future work
There is another way to express bitwise properties of natural number: digits 2
. The two ways
should be connected.
Keywords
bitwise, and, or, xor
theorem
nat.bitwise_comm
{f : bool → bool → bool}
(hf : ∀ (b b' : bool), f b b' = f b' b)
(hf' : f ff ff = ff)
(n m : ℕ) :
nat.bitwise f n m = nat.bitwise f m n
If f
is a commutative operation on bools such that f ff ff = ff
, then bitwise f
is also
commutative.
Proving associativity of bitwise operations in general essentially boils down to a huge case distinction, so it is shorter to use this tactic instead of proving it in the general case.