forked from leanprover/lean4
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathCheck.lean
122 lines (107 loc) · 3.62 KB
/
Check.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
/-
Copyright (c) 2019 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
import Lean.Meta.InferType
/-
This is not the Kernel type checker, but an auxiliary method for checking
whether terms produced by tactics and `isDefEq` are type correct.
-/
namespace Lean
namespace Meta
private def ensureType (e : Expr) : MetaM Unit := do
_ ← getLevel e; pure ()
def throwLetTypeMismatchMessage {α} (fvarId : FVarId) : MetaM α := do
lctx ← getLCtx;
match lctx.find? fvarId with
| some (LocalDecl.ldecl _ n t v b _) => do
vType ← inferType v;
throwError $
"invalid let declaration, term" ++ indentExpr v
++ Format.line ++ "has type " ++ indentExpr vType
++ Format.line ++ "but is expected to have type" ++ indentExpr t
| _ => unreachable!
@[specialize] private def checkLambdaLet
(check : Expr → MetaM Unit)
(e : Expr) : MetaM Unit :=
lambdaLetTelescope e $ fun xs b => do
xs.forM $ fun x => do {
xDecl ← getFVarLocalDecl x;
match xDecl with
| LocalDecl.cdecl _ _ _ t _ => do
ensureType t;
check t
| LocalDecl.ldecl _ _ _ t v _ => do
ensureType t;
check t;
vType ← inferType v;
unlessM (Meta.isExprDefEqAux t vType) $ throwLetTypeMismatchMessage x.fvarId!;
check v
};
check b
@[specialize] private def checkForall
(check : Expr → MetaM Unit)
(e : Expr) : MetaM Unit :=
forallTelescope e $ fun xs b => do
xs.forM $ fun x => do {
xDecl ← getFVarLocalDecl x;
ensureType xDecl.type;
check xDecl.type
};
ensureType b;
check b
private def checkConstant (constName : Name) (us : List Level) : MetaM Unit := do
cinfo ← getConstInfo constName;
unless (us.length == cinfo.lparams.length) $ throwIncorrectNumberOfLevels constName us
private def getFunctionDomain (f : Expr) : MetaM Expr := do
fType ← inferType f;
fType ← whnfD fType;
match fType with
| Expr.forallE _ d _ _ => pure d
| _ => throwFunctionExpected f
def throwAppTypeMismatch {α} (f a : Expr) (extraMsg : MessageData := Format.nil) : MetaM α := do
let e := mkApp f a;
aType ← inferType a;
expectedType ← getFunctionDomain f;
throwError $
"application type mismatch" ++ indentExpr e
++ Format.line ++ "argument" ++ indentExpr a
++ Format.line ++ "has type" ++ indentExpr aType
++ Format.line ++ "but is expected to have type" ++ indentExpr expectedType
++ extraMsg
@[specialize] private def checkApp
(check : Expr → MetaM Unit)
(f a : Expr) : MetaM Unit := do
check f;
check a;
fType ← inferType f;
fType ← whnf fType;
match fType with
| Expr.forallE _ d _ _ => do
aType ← inferType a;
unlessM (Meta.isExprDefEqAux d aType) $ throwAppTypeMismatch f a
| _ => throwFunctionExpected (mkApp f a)
private partial def checkAux : Expr → MetaM Unit
| e@(Expr.forallE _ _ _ _) => checkForall checkAux e
| e@(Expr.lam _ _ _ _) => checkLambdaLet checkAux e
| e@(Expr.letE _ _ _ _ _) => checkLambdaLet checkAux e
| Expr.const c lvls _ => checkConstant c lvls
| Expr.app f a _ => checkApp checkAux f a
| Expr.mdata _ e _ => checkAux e
| Expr.proj _ _ e _ => checkAux e
| _ => pure ()
def check (e : Expr) : MetaM Unit :=
traceCtx `Meta.check $
withTransparency TransparencyMode.all $ checkAux e
def isTypeCorrect (e : Expr) : MetaM Bool :=
catch
(traceCtx `Meta.check $ do checkAux e; pure true)
(fun ex => do
trace! `Meta.typeError ex.toMessageData;
pure false)
@[init] private def regTraceClasses : IO Unit := do
registerTraceClass `Meta.check;
registerTraceClass `Meta.typeError
end Meta
end Lean