Skip to content
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
69 changes: 63 additions & 6 deletions src/Init/Data/Bool.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,6 +12,57 @@ public section

namespace Bool

/--
Boolean complement, or “logical not”. `lnot x`.

`lnot x` is `true` when `x` is `false`. It is functionally the same as
`!x` but it uses non-branching code at runtime.

Examples:
* `lnot true = false`
* `lnot false = true`
-/
@[expose, extern "lean_bool_complement"]
def lnot (x : Bool) : Bool := !x

@[simp] theorem lnot_eq_not (x : Bool) : lnot x = (!x) := rfl

/--
Boolean “logical and”. `land x y`.

`land x y` is `true` when both of `x` or `y` are `true`. It is functionally the same as
`x && y` but it does not have short-circuiting behavior: any call to `land` will evaluate both
arguments. It uses non-branching code at runtime.

Examples:
* `land false false = false`
* `land true false = false`
* `land false true = false`
* `land true true = true`
-/
@[expose, extern "lean_bool_land"]
def land (x y : Bool) : Bool := x && y
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Is there a name that makes it easier to understand the difference between this and Bool.and at a glance? Maybe something like strictAnd?

Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

strictAnd already exists. It ensures that both arguments get evaluated but still has a branching implementation.

Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Maybe bitAnd? I can't think of anything other than land that is commonly used for this.

Copy link
Member

@TwoFX TwoFX Nov 10, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Is there a reason to have both strictAnd and Bool.land? Would strictAnd just be a worse version of Bool.land?


@[simp] theorem land_eq_and (x y : Bool) : land x y = (x && y) := rfl

/--
Boolean “logical or”. `lor x y`.

`lor x y` is `true` when at least one of `x` or `y` is `true`. It is functionally the same as
`x || y` but it does not have short-circuiting behavior: any call to `lor` will evaluate both
arguments. It uses non-branching code at runtime.

Examples:
* `lor false false = false`
* `lor true false = true`
* `lor false true = true`
* `lor true true = true`
-/
@[expose, extern "lean_bool_lor"]
def lor (x y : Bool) : Bool := x || y

@[simp] theorem lor_eq_or (x y : Bool) : lor x y = (x || y) := rfl

/--
Boolean “exclusive or”. `xor x y` can be written `x ^^ y`.

Expand All @@ -25,11 +76,14 @@ Examples:
* `false ^^ true = true`
* `true ^^ true = false`
-/
abbrev xor : Bool → Bool → Bool := bne
@[expose, reducible, extern "lean_bool_xor"]
def xor (x y : Bool) : Bool := bne x y

@[inherit_doc] infixl:33 " ^^ " => Bool.xor

@[inherit_doc] infixl:33 " ^^ " => xor
recommended_spelling "xor" for "^^" in [Bool.xor, «term_^^_»]

recommended_spelling "xor" for "^^" in [xor, «term_^^_»]
theorem xor_eq_bne (x y : Bool) : (x ^^ y) = (x != y) := rfl

instance (p : Bool → Prop) [inst : DecidablePred p] : Decidable (∀ x, p x) :=
match inst true, inst false with
Expand All @@ -48,11 +102,13 @@ instance (p : Bool → Prop) [inst : DecidablePred p] : Decidable (∃ x, p x) :
instance : LE Bool := ⟨(. → .)⟩
instance : LT Bool := ⟨(!. && .)⟩

@[extern "lean_bool_dec_le"]
instance (x y : Bool) : Decidable (x ≤ y) := inferInstanceAs (Decidable (x → y))
@[extern "lean_bool_dec_lt"]
instance (x y : Bool) : Decidable (x < y) := inferInstanceAs (Decidable (!x && y))

instance : Max Bool := ⟨or
instance : Min Bool := ⟨and
instance : Max Bool := ⟨lor
instance : Min Bool := ⟨land

theorem false_ne_true : false ≠ true := Bool.noConfusion

Expand Down Expand Up @@ -385,7 +441,8 @@ theorem and_or_inj_left_iff :
/--
Converts `true` to `1` and `false` to `0`.
-/
@[expose] def toNat (b : Bool) : Nat := cond b 1 0
@[expose, extern "lean_bool_to_nat"]
def toNat (b : Bool) : Nat := cond b 1 0

@[simp, bitvec_to_nat, grind =] theorem toNat_false : false.toNat = 0 := rfl

Expand Down
9 changes: 8 additions & 1 deletion src/include/lean/lean.h
Original file line number Diff line number Diff line change
Expand Up @@ -1780,6 +1780,7 @@ static inline uint8_t lean_int_dec_nonneg(b_lean_obj_arg a) {

/* Bool */

static inline lean_obj_res lean_bool_to_nat(uint8_t a) { return lean_usize_to_nat((size_t)a); }
static inline uint8_t lean_bool_to_uint8(uint8_t a) { return a; }
static inline uint16_t lean_bool_to_uint16(uint8_t a) { return (uint16_t)a; }
static inline uint32_t lean_bool_to_uint32(uint8_t a) { return (uint32_t)a; }
Expand All @@ -1790,7 +1791,13 @@ static inline uint16_t lean_bool_to_int16(uint8_t a) { return (uint16_t)(int16_t
static inline uint32_t lean_bool_to_int32(uint8_t a) { return (uint32_t)(int32_t)a; }
static inline uint64_t lean_bool_to_int64(uint8_t a) { return (uint64_t)(int64_t)a; }
static inline size_t lean_bool_to_isize(uint8_t a) { return (size_t)(ptrdiff_t)a; }

static inline uint8_t lean_bool_land(uint8_t a, uint8_t b) { return a & b; }
static inline uint8_t lean_bool_lor(uint8_t a, uint8_t b) { return a | b; }
static inline uint8_t lean_bool_xor(uint8_t a, uint8_t b) { return a ^ b; }
static inline uint8_t lean_bool_complement(uint8_t a) { return a ^ 1; }
static inline uint8_t lean_bool_dec_eq(uint8_t a, uint8_t b) { return a == b; }
static inline uint8_t lean_bool_dec_lt(uint8_t a, uint8_t b) { return a < b; }
static inline uint8_t lean_bool_dec_le(uint8_t a, uint8_t b) { return a <= b; }

/* UInt8 */

Expand Down
Loading