Skip to content

Commit 8897bd8

Browse files
committed
fix: teach Exception.isRuntime to detect nested errors
1 parent 5bd331e commit 8897bd8

File tree

4 files changed

+24
-11
lines changed

4 files changed

+24
-11
lines changed

src/Lean/CoreM.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -647,7 +647,10 @@ export Core (CoreM mkFreshUserName checkSystem withCurrHeartbeats)
647647
This function is a bit hackish. The heartbeat exception should probably be an internal exception.
648648
We used a similar hack at `Exception.isMaxRecDepth` -/
649649
def Exception.isMaxHeartbeat (ex : Exception) : Bool :=
650-
ex matches Exception.error _ (.tagged `runtime.maxHeartbeats _)
650+
if let Exception.error _ msg := ex then
651+
msg.stripNestedTags.kind == `runtime.maxHeartbeats
652+
else
653+
false
651654

652655
/-- Creates the expression `d → b` -/
653656
def mkArrow (d b : Expr) : CoreM Expr :=

src/Lean/Exception.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -232,7 +232,10 @@ but it is also produced by `MacroM` which implemented in the prelude, and intern
232232
been defined yet.
233233
-/
234234
def Exception.isMaxRecDepth (ex : Exception) : Bool :=
235-
ex matches error _ (.tagged `runtime.maxRecDepth _)
235+
if let Exception.error _ msg := ex then
236+
msg.stripNestedTags.kind == `runtime.maxRecDepth
237+
else
238+
false
236239

237240
/--
238241
Increment the current recursion depth and then execute `x`.

src/Lean/Log.lean

Lines changed: 1 addition & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -86,8 +86,7 @@ If `msg` is tagged as a named error, appends the error description widget displa
8686
corresponding error name and explanation link. Otherwise, returns `msg` unaltered.
8787
-/
8888
private def MessageData.appendDescriptionWidgetIfNamed (msg : MessageData) : MessageData :=
89-
let kind := stripNestedTags msg.kind
90-
match errorNameOfKind? kind with
89+
match msg.stripNestedTags.errorName? with
9190
| some errorName =>
9291
let url := manualRoot ++ s!"find/?domain={errorExplanationManualDomain}&name={errorName}"
9392
let inst := {
@@ -102,13 +101,6 @@ private def MessageData.appendDescriptionWidgetIfNamed (msg : MessageData) : Mes
102101
-- console output
103102
msg.composePreservingKind <| .ofWidget inst .nil
104103
| none => msg
105-
where
106-
/-- Remove any `` `nested `` name components prepended by `throwNestedTacticEx`. -/
107-
stripNestedTags : Name → Name
108-
| .str p "nested" => stripNestedTags p
109-
| .str p s => .str (stripNestedTags p) s
110-
| .num p n => .num (stripNestedTags p) n
111-
| .anonymous => .anonymous
112104

113105
/--
114106
Log the message `msgData` at the position provided by `ref` with the given `severity`.

src/Lean/Message.lean

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -457,6 +457,21 @@ and `throwNamedError`.
457457
def MessageData.tagWithErrorName (msg : MessageData) (name : Name) : MessageData :=
458458
.tagged (kindOfErrorName name) msg
459459

460+
/-- Strip the `` `nested`` prefix components added to tags by `throwNestedTacticEx`. -/
461+
def MessageData.stripNestedTags : MessageData → MessageData
462+
| .tagged n msg => .tagged (stripNestedNamePrefix n) msg
463+
| msg => msg
464+
where
465+
stripNestedNamePrefix : Name → Name
466+
| .anonymous => .anonymous
467+
| .str p s =>
468+
let p' := stripNestedNamePrefix p
469+
if p'.isAnonymous && s == "nested" then
470+
.anonymous
471+
else
472+
.str p' s
473+
| .num p n => .num (stripNestedNamePrefix p) n
474+
460475
/--
461476
If the provided name is labeled as a diagnostic name, removes the label and returns the
462477
corresponding diagnostic name.

0 commit comments

Comments
 (0)