diff --git a/src/Lean/Log.lean b/src/Lean/Log.lean index d3a09b3ad563..646c0140a337 100644 --- a/src/Lean/Log.lean +++ b/src/Lean/Log.lean @@ -60,7 +60,10 @@ We use the `fileMap` to find the line and column numbers for the error message. -/ def logAt (ref : Syntax) (msgData : MessageData) (severity : MessageSeverity := MessageSeverity.error) (isSilent : Bool := false) : m Unit := - unless severity == .error && msgData.hasSyntheticSorry do + -- Filter out follow-up errors on synthetic sorries, which should only be created when an error + -- for the relevant part of the input has already been logged. However, in case this invariant + -- was accidentally broken and no error has been logged yet, log at least one error. + unless severity == .error && msgData.hasSyntheticSorry && (← MonadLog.hasErrors) do let severity := if severity == .warning && warningAsError.get (← getOptions) then .error else severity let ref := replaceRef ref (← MonadLog.getRef) let pos := ref.getPos?.getD 0 diff --git a/tests/lean/doLetLoop.lean.expected.out b/tests/lean/doLetLoop.lean.expected.out index d7d4ca528253..7e2a9b4f6ae0 100644 --- a/tests/lean/doLetLoop.lean.expected.out +++ b/tests/lean/doLetLoop.lean.expected.out @@ -1,2 +1,5 @@ doLetLoop.lean:4:0: error: unexpected end of input -doLetLoop.lean:2:4-2:5: warning: declaration uses 'sorry' +doLetLoop.lean:3:2-3:8: error: failed to synthesize + Decidable sorry + +Additional diagnostic information may be available using the `set_option diagnostics true` command. diff --git a/tests/lean/doSeqRightIssue.lean.expected.out b/tests/lean/doSeqRightIssue.lean.expected.out index 405f86ef10e3..7050d3f58add 100644 --- a/tests/lean/doSeqRightIssue.lean.expected.out +++ b/tests/lean/doSeqRightIssue.lean.expected.out @@ -1,3 +1,7 @@ doSeqRightIssue.lean:5:23-5:24: error: unknown universe level 'v' -doSeqRightIssue.lean:8:0-9:40: warning: declaration uses 'sorry' -doSeqRightIssue.lean:7:8-7:10: warning: declaration uses 'sorry' +doSeqRightIssue.lean:7:36-7:39: error: type mismatch + β a +has type + sorry : Sort ?u +but is expected to have type + Type ?u : Type (?u + 1) diff --git a/tests/lean/run/2226.lean b/tests/lean/run/2226.lean index 810f37c18eef..f252618d8626 100644 --- a/tests/lean/run/2226.lean +++ b/tests/lean/run/2226.lean @@ -6,7 +6,14 @@ A : Nat #guard_msgs in variable (A : Nat) (B : by skip) -/-- error: failed to infer definition type -/ +/-- +error: type mismatch + B +has type + sorry : Sort ?u.9 +but is expected to have type + Nat : Type +-/ #guard_msgs in def foo := A = B diff --git a/tests/lean/run/3214.lean b/tests/lean/run/3214.lean index 0c7f11a6d10d..71558984c695 100644 --- a/tests/lean/run/3214.lean +++ b/tests/lean/run/3214.lean @@ -11,4 +11,15 @@ term has type #guard_msgs in variable {α : Type} (s : Missing α) +/-- +error: application type mismatch + Foo s +argument + s +has type + sorry : Sort _ +but is expected to have type + Type : Type 1 +-/ +#guard_msgs in #synth Foo s diff --git a/tests/lean/silent_failure.lean b/tests/lean/silent_failure.lean new file mode 100644 index 000000000000..7fc0ac77baaa --- /dev/null +++ b/tests/lean/silent_failure.lean @@ -0,0 +1,7 @@ +-- Prior to https://github.com/leanprover/lean4/pull/8230 +-- this failing silently (but not adding `g` to the environment). + +#guard_msgs (drop info) in +partial def g : + have : False := by apply? + False := g diff --git a/tests/lean/silent_failure.lean.expected.out b/tests/lean/silent_failure.lean.expected.out new file mode 100644 index 000000000000..0ae8cc2807cd --- /dev/null +++ b/tests/lean/silent_failure.lean.expected.out @@ -0,0 +1,3 @@ +silent_failure.lean:5:12-5:13: error: invalid use of 'partial', 'g' is not a function + let_fun this := ⋯; + False