Skip to content

Commit

Permalink
fix: collect level parameters in evalExpr
Browse files Browse the repository at this point in the history
`elabEvalUnsafe` already does something similar.
  • Loading branch information
eric-wieser committed Dec 19, 2023
1 parent 62c3e56 commit 8335d4e
Show file tree
Hide file tree
Showing 2 changed files with 16 additions and 1 deletion.
4 changes: 3 additions & 1 deletion src/Lean/Meta/Eval.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,19 +4,21 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Sebastian Ullrich, Leonardo de Moura
-/
import Lean.Meta.Check
import Lean.Util.CollectLevelParams

namespace Lean.Meta

unsafe def evalExprCore (α) (value : Expr) (checkType : Expr → MetaM Unit) (safety := DefinitionSafety.safe) : MetaM α :=
withoutModifyingEnv do
let name ← mkFreshUserName `_tmp
let value ← instantiateMVars value
let us := collectLevelParams {} value |>.params
if value.hasMVar then
throwError "failed to evaluate expression, it contains metavariables{indentExpr value}"
let type ← inferType value
checkType type
let decl := Declaration.defnDecl {
name, levelParams := [], type
name, levelParams := us.toList, type
value, hints := ReducibilityHints.opaque,
safety
}
Expand Down
13 changes: 13 additions & 0 deletions tests/lean/run/3091.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
import Lean

open Lean

def foo.{u} : Nat := (ULift.up.{u} Nat.zero).down

universe u in
#eval foo.{u}

-- this used to produce an error
#eval do
let u : Lean.Level := .param `u
Meta.evalExpr Nat (.const ``Nat []) (.const ``foo [u])

0 comments on commit 8335d4e

Please sign in to comment.