Skip to content

Commit

Permalink
fix: Option.getD eagerly evaluates dflt (#3043)
Browse files Browse the repository at this point in the history
Reported [on
Zulip](https://leanprover.zulipchat.com/#narrow/stream/348111-std4/topic/Panics.20in.20Std.2EHashMap.2Efind!/near/406872395).
The `dflt` argument of `Option.getD` is not evaluated lazily, as the
documentation says, because even after `macro_inline` the expression
```lean
match opt, dflt with
| some x, _ => x
| none, e => e
```
still has the semantics of evaluating `dflt` when `opt` is `some x`.
  • Loading branch information
digama0 authored Dec 11, 2023
1 parent e6c0484 commit 178ab8e
Show file tree
Hide file tree
Showing 4 changed files with 20 additions and 6 deletions.
7 changes: 4 additions & 3 deletions src/Init/Prelude.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2213,9 +2213,10 @@ returns `a` if `opt = some a` and `dflt` otherwise.
This function is `@[macro_inline]`, so `dflt` will not be evaluated unless
`opt` turns out to be `none`.
-/
@[macro_inline] def Option.getD : Option α → α → α
| some x, _ => x
| none, e => e
@[macro_inline] def Option.getD (opt : Option α) (dflt : α) : α :=
match opt with
| some x => x
| none => dflt

/--
Map a function over an `Option` by applying the function to the contained
Expand Down
12 changes: 12 additions & 0 deletions tests/lean/optionGetD.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@
import Lean.Data.HashMap

def test (m : Lean.HashMap Nat Nat) : IO (Nat × Nat) := do
let start := 1
let mut i := start
let mut count := 0
while i != 0 do
i := (m.find? i).getD (panic! "key is not in the map")
count := count + 1
return (i, count)

#eval test (.ofList [(1,3),(3,2),(2,0)])
1 change: 1 addition & 0 deletions tests/lean/optionGetD.lean.expected.out
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
(0, 3)
6 changes: 3 additions & 3 deletions tests/lean/smartUnfoldingMatch.lean.expected.out
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
42
[Meta.debug] r: match x, 0 with
| some x, x_1 => x
| none, e => e
[Meta.debug] r: match x with
| some x => x
| none => 0

0 comments on commit 178ab8e

Please sign in to comment.