forked from leanprover/lean4
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathCheck.lean
202 lines (184 loc) · 6.9 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
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
/-
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.Meta
private def ensureType (e : Expr) : MetaM Unit := do
discard <| getLevel e
def throwLetTypeMismatchMessage {α} (fvarId : FVarId) : MetaM α := do
let lctx ← getLCtx
match lctx.find? fvarId with
| some (LocalDecl.ldecl _ _ _ t v _) => do
let vType ← inferType v
throwError "invalid let declaration, term{indentExpr v}\nhas type{indentExpr vType}\nbut is expected to have type{indentExpr t}"
| _ => unreachable!
private def checkConstant (constName : Name) (us : List Level) : MetaM Unit := do
let cinfo ← getConstInfo constName
unless us.length == cinfo.levelParams.length do
throwIncorrectNumberOfLevels constName us
private def getFunctionDomain (f : Expr) : MetaM (Expr × BinderInfo) := do
let fType ← inferType f
let fType ← whnfD fType
match fType with
| Expr.forallE _ d _ c => return (d, c)
| _ => throwFunctionExpected f
/-
Given to expressions `a` and `b`, this method tries to annotate terms with `pp.explicit := true` to
expose "implicit" differences. For example, suppose `a` and `b` are of the form
```lean
@HashMap Nat Nat eqInst hasInst1
@HashMap Nat Nat eqInst hasInst2
```
By default, the pretty printer formats both of them as `HashMap Nat Nat`.
So, counterintuitive error messages such as
```lean
error: application type mismatch
HashMap.insert m
argument
m
has type
HashMap Nat Nat
but is expected to have type
HashMap Nat Nat
```
would be produced.
By adding `pp.explicit := true`, we can generate the more informative error
```lean
error: application type mismatch
HashMap.insert m
argument
m
has type
@HashMap Nat Nat eqInst hasInst1
but is expected to have type
@HashMap Nat Nat eqInst hasInst2
```
Remark: this method implements a simple heuristic, we should extend it as we find other counterintuitive
error messages.
-/
partial def addPPExplicitToExposeDiff (a b : Expr) : MetaM (Expr × Expr) := do
if (← getOptions).getBool `pp.all false || (← getOptions).getBool `pp.explicit false then
return (a, b)
else
visit (← instantiateMVars a) (← instantiateMVars b)
where
visit (a b : Expr) : MetaM (Expr × Expr) := do
try
if !a.isApp || !b.isApp then
return (a, b)
else if a.getAppNumArgs != b.getAppNumArgs then
return (a, b)
else if not (← isDefEq a.getAppFn b.getAppFn) then
return (a, b)
else
let fType ← inferType a.getAppFn
forallBoundedTelescope fType a.getAppNumArgs fun xs _ => do
let mut as := a.getAppArgs
let mut bs := b.getAppArgs
if let some (as', bs') ← hasExplicitDiff? xs as bs then
return (mkAppN a.getAppFn as', mkAppN b.getAppFn bs')
else
for i in [:as.size] do
unless (← isDefEq as[i]! bs[i]!) do
let (ai, bi) ← visit as[i]! bs[i]!
as := as.set! i ai
bs := bs.set! i bi
let a := mkAppN a.getAppFn as
let b := mkAppN b.getAppFn bs
return (a.setAppPPExplicit, b.setAppPPExplicit)
catch _ =>
return (a, b)
hasExplicitDiff? (xs as bs : Array Expr) : MetaM (Option (Array Expr × Array Expr)) := do
for i in [:xs.size] do
let localDecl ← getLocalDecl xs[i]!.fvarId!
if localDecl.binderInfo.isExplicit then
unless (← isDefEq as[i]! bs[i]!) do
let (ai, bi) ← visit as[i]! bs[i]!
return some (as.set! i ai, bs.set! i bi)
return none
/-
Return error message "has type{givenType}\nbut is expected to have type{expectedType}"
-/
def mkHasTypeButIsExpectedMsg (givenType expectedType : Expr) : MetaM MessageData := do
try
let givenTypeType ← inferType givenType
let expectedTypeType ← inferType expectedType
let (givenType, expectedType) ← addPPExplicitToExposeDiff givenType expectedType
let (givenTypeType, expectedTypeType) ← addPPExplicitToExposeDiff givenTypeType expectedTypeType
return m!"has type{indentD m!"{givenType} : {givenTypeType}"}\nbut is expected to have type{indentD m!"{expectedType} : {expectedTypeType}"}"
catch _ =>
let (givenType, expectedType) ← addPPExplicitToExposeDiff givenType expectedType
return m!"has type{indentExpr givenType}\nbut is expected to have type{indentExpr expectedType}"
def throwAppTypeMismatch (f a : Expr) : MetaM α := do
let (expectedType, binfo) ← getFunctionDomain f
let mut e := mkApp f a
unless binfo.isExplicit do
e := e.setAppPPExplicit
let aType ← inferType a
throwError "application type mismatch{indentExpr e}\nargument{indentExpr a}\n{← mkHasTypeButIsExpectedMsg aType expectedType}"
def checkApp (f a : Expr) : MetaM Unit := do
let fType ← inferType f
let fType ← whnf fType
match fType with
| Expr.forallE _ d _ _ =>
let aType ← inferType a
unless (← isDefEq d aType) do
throwAppTypeMismatch f a
| _ => throwFunctionExpected (mkApp f a)
private partial def checkAux (e : Expr) : MetaM Unit := do
check e |>.run
where
check (e : Expr) : MonadCacheT ExprStructEq Unit MetaM Unit :=
checkCache { val := e : ExprStructEq } fun _ => do
match e with
| .forallE .. => checkForall e
| .lam .. => checkLambdaLet e
| .letE .. => checkLambdaLet e
| .const c lvls => checkConstant c lvls
| .app f a => check f; check a; checkApp f a
| .mdata _ e => check e
| .proj _ _ e => check e
| _ => return ()
checkLambdaLet (e : Expr) : MonadCacheT ExprStructEq Unit MetaM Unit :=
lambdaLetTelescope e fun xs b => do
xs.forM fun x => do
let xDecl ← getFVarLocalDecl x;
match xDecl with
| LocalDecl.cdecl (type := t) .. =>
ensureType t
check t
| LocalDecl.ldecl (type := t) (value := v) .. =>
ensureType t
check t
let vType ← inferType v
unless (← isDefEq t vType) do throwLetTypeMismatchMessage x.fvarId!
check v
check b
checkForall (e : Expr) : MonadCacheT ExprStructEq Unit MetaM Unit :=
forallTelescope e fun xs b => do
xs.forM fun x => do
let xDecl ← getFVarLocalDecl x
ensureType xDecl.type
check xDecl.type
ensureType b
check b
def check (e : Expr) : MetaM Unit :=
traceCtx `Meta.check do
withTransparency TransparencyMode.all $ checkAux e
def isTypeCorrect (e : Expr) : MetaM Bool := do
try
check e
pure true
catch ex =>
trace[Meta.typeError] ex.toMessageData
pure false
builtin_initialize
registerTraceClass `Meta.check
registerTraceClass `Meta.typeError
end Lean.Meta