Skip to content

Commit

Permalink
fix: check must check projections
Browse files Browse the repository at this point in the history
This PR ensures `Meta.check` check projections.

closes #5660
  • Loading branch information
leodemoura committed Dec 16, 2024
1 parent a8a160b commit a7d132f
Show file tree
Hide file tree
Showing 3 changed files with 74 additions and 5 deletions.
8 changes: 7 additions & 1 deletion src/Lean/Meta/Check.lean
Original file line number Diff line number Diff line change
Expand Up @@ -209,6 +209,12 @@ def checkApp (f a : Expr) : MetaM Unit := do
throwAppTypeMismatch f a
| _ => throwFunctionExpected (mkApp f a)

def checkProj (structName : Name) (idx : Nat) (e : Expr) : MetaM Unit := do
let structType ← whnf (← inferType e)
let projType ← inferType (mkProj structName idx e)
if (← isProp structType) && !(← isProp projType) then
throwError "invalid projection{indentExpr (mkProj structName idx e)}\nfrom type{indentExpr structType}"

private partial def checkAux (e : Expr) : MetaM Unit := do
check e |>.run
where
Expand All @@ -221,7 +227,7 @@ where
| .const c lvls => checkConstant c lvls
| .app f a => check f; check a; checkApp f a
| .mdata _ e => check e
| .proj _ _ e => check e
| .proj s i e => check e; checkProj s i e
| _ => return ()

checkLambdaLet (e : Expr) : MonadCacheT ExprStructEq Unit MetaM Unit :=
Expand Down
10 changes: 6 additions & 4 deletions src/Lean/Meta/InferType.lean
Original file line number Diff line number Diff line change
Expand Up @@ -99,7 +99,9 @@ private def inferProjType (structName : Name) (idx : Nat) (e : Expr) : MetaM Exp
let structType ← whnf structType
let failed {α} : Unit → MetaM α := fun _ => do
throwError "invalid projection{indentExpr (mkProj structName idx e)}\nfrom type{indentExpr structType}"
matchConstStructure structType.getAppFn failed fun structVal structLvls ctorVal =>
matchConstStructure structType.getAppFn failed fun structVal structLvls ctorVal => do
unless structVal.name == structName do
failed ()
let structTypeArgs := structType.getAppArgs
if structVal.numParams + structVal.numIndices != structTypeArgs.size then
failed ()
Expand All @@ -108,16 +110,16 @@ private def inferProjType (structName : Name) (idx : Nat) (e : Expr) : MetaM Exp
for i in [:idx] do
ctorType ← whnf ctorType
match ctorType with
| Expr.forallE _ _ body _ =>
| .forallE _ _ body _ =>
if body.hasLooseBVars then
ctorType := body.instantiate1 <| mkProj structName i e
else
ctorType := body
| _ => failed ()
ctorType ← whnf ctorType
match ctorType with
| Expr.forallE _ d _ _ => return d.consumeTypeAnnotations
| _ => failed ()
| .forallE _ d _ _ => return d.consumeTypeAnnotations
| _ => failed ()

def throwTypeExcepted {α} (type : Expr) : MetaM α :=
throwError "type expected{indentExpr type}"
Expand Down
61 changes: 61 additions & 0 deletions tests/lean/run/5660.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,61 @@
import Lean


def ex : ∃ x : Nat, x = 1 :=
1, rfl⟩

open Lean Meta

/--
error: invalid projection
ex.3
from type
∃ x, x = 1
-/
#guard_msgs in
run_meta check (Expr.proj ``Exists 2 (mkConst ``ex)) -- should fail

/--
error: invalid projection
ex.1
from type
∃ x, x = 1
-/
#guard_msgs in
run_meta check (Expr.proj ``Exists 0 (mkConst ``ex)) -- should fail

run_meta check (Expr.proj ``Exists 1 (mkConst ``ex))

def p := (1, 2)

/--
error: invalid projection
p.1
from type
Nat × Nat
-/
#guard_msgs in
run_meta check (Expr.proj ``Exists 0 (mkConst ``p)) -- should fail

run_meta check (Expr.proj ``Prod 0 (mkConst ``p))
run_meta check (Expr.proj ``Prod 1 (mkConst ``p))

/--
error: invalid projection
p.3
from type
Nat × Nat
-/
#guard_msgs in
run_meta check (Expr.proj ``Prod 2 (mkConst ``p)) -- should fail

def n := 1

/--
error: invalid projection
n.1
from type
Nat
-/
#guard_msgs in
run_meta check (Expr.proj ``Nat 0 (mkConst ``n)) -- should fail

0 comments on commit a7d132f

Please sign in to comment.