Skip to content

Commit

Permalink
chore: make List.all and List.any short-circuit
Browse files Browse the repository at this point in the history
@[inline] doesn't do anything now?
  • Loading branch information
semorrison committed Nov 28, 2023
1 parent 7ff7cf9 commit 8009ae5
Showing 1 changed file with 10 additions and 4 deletions.
14 changes: 10 additions & 4 deletions src/Init/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -557,16 +557,22 @@ def takeWhile (p : α → Bool) : (xs : List α) → List α
/--
`O(|l|)`. Returns true if `p` is true for any element of `l`.
* `any p [a, b, c] = p a || p b || p c`
Short-circuits upon encountering the first `true`.
-/
@[inline] def any (l : List α) (p : α → Bool) : Bool :=
foldr (fun a r => p a || r) false l
def any : List α -> (α → Bool) -> Bool
| [], _ => false
| h :: t, p => p h || any t p

/--
`O(|l|)`. Returns true if `p` is true for every element of `l`.
* `all p [a, b, c] = p a && p b && p c`
Short-circuits upon encountering the first `false`.
-/
@[inline] def all (l : List α) (p : α → Bool) : Bool :=
foldr (fun a r => p a && r) true l
def all : List α -> (α → Bool) -> Bool
| [], _ => true
| h :: t, p => p h && all t p

/--
`O(|l|)`. Returns true if `true` is an element of the list of booleans `l`.
Expand Down

0 comments on commit 8009ae5

Please sign in to comment.