Skip to content

Commit 0a44c68

Browse files
committed
feat: expose Kernel.check for debugging purposes
along `Kernel.isDefEq` and `Kernel.whnf`.
1 parent a6830f9 commit 0a44c68

File tree

3 files changed

+23
-0
lines changed

3 files changed

+23
-0
lines changed

src/Lean/Environment.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1096,6 +1096,13 @@ def isDefEqGuarded (env : Environment) (lctx : LocalContext) (a b : Expr) : Bool
10961096
@[extern "lean_kernel_whnf"]
10971097
opaque whnf (env : Environment) (lctx : LocalContext) (a : Expr) : Except KernelException Expr
10981098

1099+
/--
1100+
Kernel typecheck function. We use it mainly for debugging purposes.
1101+
Recall that the Kernel type checker does not support metavariables.
1102+
When implementing automation, consider using the `MetaM` methods. -/
1103+
@[extern "lean_kernel_check"]
1104+
opaque check (env : Environment) (lctx : LocalContext) (a : Expr) : Except KernelException Expr
1105+
10991106
end Kernel
11001107

11011108
class MonadEnv (m : TypeType) where

src/kernel/type_checker.cpp

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1203,6 +1203,12 @@ extern "C" LEAN_EXPORT lean_object * lean_kernel_whnf(lean_object * env, lean_ob
12031203
});
12041204
}
12051205

1206+
extern "C" LEAN_EXPORT lean_object * lean_kernel_check(lean_object * env, lean_object * lctx, lean_object * a) {
1207+
return catch_kernel_exceptions<object*>([&]() {
1208+
return type_checker(environment(env), local_ctx(lctx)).check(expr(a)).steal();
1209+
});
1210+
}
1211+
12061212
inline static expr * new_persistent_expr_const(name const & n) {
12071213
expr * e = new expr(mk_const(n));
12081214
mark_persistent(e->raw());

tests/lean/run/kernel2.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -15,6 +15,12 @@ def whnf (a : Name) : CoreM Unit := do
1515
let r ← ofExceptKernelException (Kernel.whnf env {} a)
1616
IO.println (toString a ++ " ==> " ++ toString r)
1717

18+
def check (a : Name) : CoreM Unit := do
19+
let env ← getEnv
20+
let a := mkConst a
21+
let r ← ofExceptKernelException (Kernel.check env {} a)
22+
IO.println (toString a ++ " ==> " ++ toString r)
23+
1824
partial def fact : Nat → Nat
1925
| 0 => 1
2026
| (n+1) => (n+1)*fact n
@@ -93,3 +99,7 @@ def c12 : Nat := 'a'.toNat
9399
/-- info: c12 ==> 97 -/
94100
#guard_msgs in
95101
#eval whnf `c12
102+
103+
/-- info: c11 ==> Bool -/
104+
#guard_msgs in
105+
#eval check `c11

0 commit comments

Comments
 (0)