Skip to content

Commit

Permalink
feat: have IR checker suggest noncomputable (#4729)
Browse files Browse the repository at this point in the history
Currently, `ll_infer_type` is responsible for telling the user about
`noncomputable` when a definition depends on one without executable
code. However, this is imperfect because type inference does not check
every subexpression. This leads to errors later on that users find to be
hard to interpret.

Now, `Lean.IR.checkDecls` has a friendlier error message when it
encounters constants without compiled definitions, suggesting to
consider using `noncomputable`. While this function is an internal IR
consistency check, it is also reasonable to have it give an informative
error message in this particular case. The suggestion to use
`noncomputable` is limited to just unknown constants.

Some alternatives would be to either (1) create another checker just for
missing constants, (2) change `ll_infer_type` to always visit every
subexpression no matter if they are necessary for inferring the type, or
(3) investigate whether `tests/lean/run/1785.lean` is due to a deeper
issue.

Closes #1785
  • Loading branch information
kmill authored Sep 7, 2024
1 parent c9239bf commit 7a7440f
Show file tree
Hide file tree
Showing 2 changed files with 23 additions and 2 deletions.
4 changes: 2 additions & 2 deletions src/Lean/Compiler/IR/Checker.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,7 +50,7 @@ def markJP (j : JoinPointId) : M Unit :=
def getDecl (c : Name) : M Decl := do
let ctx ← read
match findEnvDecl' ctx.env c ctx.decls with
| none => throw s!"unknown declaration '{c}'"
| none => throw s!"depends on declaration '{c}', which has no executable code; consider marking definition as 'noncomputable'"
| some d => pure d

def checkVar (x : VarId) : M Unit := do
Expand Down Expand Up @@ -182,7 +182,7 @@ end Checker
def checkDecl (decls : Array Decl) (decl : Decl) : CompilerM Unit := do
let env ← getEnv
match (Checker.checkDecl decl { env := env, decls := decls }).run' {} with
| .error msg => throw s!"compiler IR check failed at '{decl.name}', error: {msg}"
| .error msg => throw s!"failed to compile definition, compiler IR check failed at '{decl.name}'. Error: {msg}"
| _ => pure ()

def checkDecls (decls : Array Decl) : CompilerM Unit :=
Expand Down
21 changes: 21 additions & 0 deletions tests/lean/run/1785.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,21 @@
/-!
# Improve compiler IR check message for users when constants are not compiled
-/

/-
This is a simplified version of the example in #1785.
Note that the error changes if the typeclass argument is removed.
-/

noncomputable
def char (R : Type) [∀ n, OfNat R n] : Nat := 0

/--
error: failed to compile definition, compiler IR check failed at 'bug._rarg'. Error: depends on
declaration 'char', which has no executable code; consider marking definition as 'noncomputable'
-/
#guard_msgs in
def bug (R : Type) [∀ n, OfNat R n] : R :=
match char R with
| 0 => 1
| _ => 0

0 comments on commit 7a7440f

Please sign in to comment.