Skip to content

Commit e5e5778

Browse files
authored
doc: mention that inferType does not ensure type correctness (#5087)
This also adds links to the implementations of `whnf` and `inferType` to make it easier to navigate this part of the code base.
1 parent 7432a6f commit e5e5778

File tree

1 file changed

+65
-9
lines changed

1 file changed

+65
-9
lines changed

src/Lean/Meta/Basic.lean

Lines changed: 65 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -560,16 +560,72 @@ def useEtaStruct (inductName : Name) : MetaM Bool := do
560560
| .all => return true
561561
| .notClasses => return !isClass (← getEnv) inductName
562562

563-
/-! WARNING: The following 4 constants are a hack for simulating forward declarations.
564-
They are defined later using the `export` attribute. This is hackish because we
565-
have to hard-code the true arity of these definitions here, and make sure the C names match.
566-
We have used another hack based on `IO.Ref`s in the past, it was safer but less efficient. -/
567-
568-
/-- Reduces an expression to its Weak Head Normal Form.
569-
This is when the topmost expression has been fully reduced,
570-
but may contain subexpressions which have not been reduced. -/
563+
/-!
564+
WARNING: The following 4 constants are a hack for simulating forward declarations.
565+
They are defined later using the `export` attribute. This is hackish because we
566+
have to hard-code the true arity of these definitions here, and make sure the C names match.
567+
We have used another hack based on `IO.Ref`s in the past, it was safer but less efficient.
568+
-/
569+
570+
/--
571+
Reduces an expression to its *weak head normal form*.
572+
This is when the "head" of the top-level expression has been fully reduced.
573+
The result may contain subexpressions that have not been reduced.
574+
575+
See `Lean.Meta.whnfImp` for the implementation.
576+
-/
571577
@[extern 6 "lean_whnf"] opaque whnf : Expr → MetaM Expr
572-
/-- Returns the inferred type of the given expression, or fails if it is not type-correct. -/
578+
/--
579+
Returns the inferred type of the given expression. Assumes the expression is type-correct.
580+
581+
The type inference algorithm does not do general type checking.
582+
Type inference only looks at subterms that are necessary for determining an expression's type,
583+
and as such if `inferType` succeeds it does *not* mean the term is type-correct.
584+
If an expression is sufficiently ill-formed that it prevents `inferType` from computing a type,
585+
then it will fail with a type error.
586+
587+
For typechecking during elaboration, see `Lean.Meta.check`.
588+
(Note that we do not guarantee that the elaborator typechecker is as correct or as efficient as
589+
the kernel typechecker. The kernel typechecker is invoked when a definition is added to the environment.)
590+
591+
Here are examples of type-incorrect terms for which `inferType` succeeds:
592+
```lean
593+
import Lean
594+
595+
open Lean Meta
596+
597+
/--
598+
`@id.{1} Bool Nat.zero`.
599+
In general, the type of `@id α x` is `α`.
600+
-/
601+
def e1 : Expr := mkApp2 (.const ``id [1]) (.const ``Bool []) (.const ``Nat.zero [])
602+
#eval inferType e1
603+
-- Lean.Expr.const `Bool []
604+
#eval check e1
605+
-- error: application type mismatch
606+
607+
/--
608+
`let x : Int := Nat.zero; true`.
609+
In general, the type of `let x := v; e`, if `e` does not reference `x`, is the type of `e`.
610+
-/
611+
def e2 : Expr := .letE `x (.const ``Int []) (.const ``Nat.zero []) (.const ``true []) false
612+
#eval inferType e2
613+
-- Lean.Expr.const `Bool []
614+
#eval check e2
615+
-- error: invalid let declaration
616+
```
617+
Here is an example of a type-incorrect term that makes `inferType` fail:
618+
```lean
619+
/--
620+
`Nat.zero Nat.zero`
621+
-/
622+
def e3 : Expr := .app (.const ``Nat.zero []) (.const ``Nat.zero [])
623+
#eval inferType e3
624+
-- error: function expected
625+
```
626+
627+
See `Lean.Meta.inferTypeImp` for the implementation of `inferType`.
628+
-/
573629
@[extern 6 "lean_infer_type"] opaque inferType : Expr → MetaM Expr
574630
@[extern 7 "lean_is_expr_def_eq"] opaque isExprDefEqAux : Expr → Expr → MetaM Bool
575631
@[extern 7 "lean_is_level_def_eq"] opaque isLevelDefEqAux : Level → Level → MetaM Bool

0 commit comments

Comments
 (0)