-
Notifications
You must be signed in to change notification settings - Fork 447
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
feat: improve error messages and docstring for
decide
tactic (#3422)
The `decide` tactic produces error messages that users find to be obscure. Now: 1. If the `Decidable` instance reduces to `isFalse`, it reports that `decide` failed because the proposition is false. 2. If the `Decidable` instance fails to reduce, it explains what proposition it failed for, and it shows the reduced `Decidable` instance rather than the `Decidable.decide` expression. That expression tends to be less useful since it shows the unreduced `Decidable` argument (plus it's a lot longer!) Examples: ```lean example : 1 ≠ 1 := by decide /- tactic 'decide' proved that the proposition 1 ≠ 1 is false -/ opaque unknownProp : Prop open scoped Classical in example : unknownProp := by decide /- tactic 'decide' failed for proposition unknownProp since its 'Decidable' instance reduced to Classical.choice ⋯ rather than to the 'isTrue' constructor. -/ ``` When reporting the error, `decide` only shows the whnf of the `Decidable` instance. In the future we could consider having it reduce all decidable instances present in the term, which can help with determining the cause of failure (this was explored in 8cede58).
- Loading branch information
Showing
6 changed files
with
113 additions
and
15 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -1,6 +1,12 @@ | ||
2161.lean:15:48-15:54: error: failed to reduce to 'true' | ||
Decidable.rec (fun h => (fun x => false) h) (fun h => (fun x => true) h) | ||
(instDecidableEqFoo (mul (mul (mul 4 1) 1) 1) 4) | ||
2161.lean:22:48-22:54: error: failed to reduce to 'true' | ||
Decidable.rec (fun h => (fun x => false) h) (fun h => (fun x => true) h) | ||
(instDecidableEqFoo (add (add (add 4 1) 1) 1) 4) | ||
2161.lean:15:48-15:54: error: tactic 'decide' failed for proposition | ||
mul (mul (mul 4 1) 1) 1 = 4 | ||
since its 'Decidable' instance reduced to | ||
Decidable.rec (fun h => (fun h => isFalse ⋯) h) (fun h => (fun h => h ▸ isTrue ⋯) h) | ||
(instDecidableEqNat (mul (mul (mul 4 1) 1) 1).num 4) | ||
rather than to the 'isTrue' constructor. | ||
2161.lean:22:48-22:54: error: tactic 'decide' failed for proposition | ||
add (add (add 4 1) 1) 1 = 4 | ||
since its 'Decidable' instance reduced to | ||
Decidable.rec (fun h => (fun h => isFalse ⋯) h) (fun h => (fun h => h ▸ isTrue ⋯) h) | ||
(instDecidableEqNat (add (add (add 4 1) 1) 1).num 4) | ||
rather than to the 'isTrue' constructor. |
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,21 @@ | ||
/-! | ||
# Tests of the `decide` tactic | ||
-/ | ||
|
||
/-! | ||
Success | ||
-/ | ||
example : 2 + 2 ≠ 5 := by decide | ||
|
||
/-! | ||
False proposition | ||
-/ | ||
example : 1 ≠ 1 := by decide | ||
|
||
/-! | ||
Irreducible decidable instance | ||
-/ | ||
opaque unknownProp : Prop | ||
|
||
open scoped Classical in | ||
example : unknownProp := by decide |
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,8 @@ | ||
decideTactic.lean:13:22-13:28: error: tactic 'decide' proved that the proposition | ||
1 ≠ 1 | ||
is false | ||
decideTactic.lean:21:28-21:34: error: tactic 'decide' failed for proposition | ||
unknownProp | ||
since its 'Decidable' instance reduced to | ||
Classical.choice ⋯ | ||
rather than to the 'isTrue' constructor. |