Skip to content

Commit

Permalink
fix: check is valid structure projection when pretty printing (#4982)
Browse files Browse the repository at this point in the history
For structure projections, the pretty printer assumed that the
expression was type correct. Now it checks that the object being
projected is of the correct type. Such terms appear in type mismatch
errors.

Also, fixes and improves `#print` for structures. The types of
projections now use MessageData (so are now hoverable), and the type of
`self` is now the correct type.

Closes #4670
  • Loading branch information
kmill authored Aug 12, 2024
1 parent ecb3579 commit 7cd406f
Show file tree
Hide file tree
Showing 3 changed files with 69 additions and 5 deletions.
10 changes: 5 additions & 5 deletions src/Lean/Elab/Print.lean
Original file line number Diff line number Diff line change
Expand Up @@ -78,19 +78,19 @@ private def printStructure (id : Name) (levelParams : List Name) (numParams : Na
logInfo m
where
doFields := liftTermElabM do
forallTelescope (← getConstInfo id).type fun params type =>
withLocalDeclD `self type fun self => do
forallTelescope (← getConstInfo id).type fun params _ =>
withLocalDeclD `self (mkAppN (Expr.const id (levelParams.map .param)) params) fun self => do
let params := params.push self
let mut m : Format := ""
let mut m : MessageData := ""
for field in fields do
match getProjFnForField? (← getEnv) id field with
| some proj =>
let field : Format := if isPrivateName proj then "private " ++ toString field else toString field
let cinfo ← getConstInfo proj
let ftype ← instantiateForall cinfo.type params
m := m ++ Format.line ++ field ++ " : " ++ (← ppExpr ftype) -- Why ppExpr here?
m := m ++ Format.line ++ field ++ " : " ++ ftype
| none => panic! "missing structure field info"
return m
addMessageContext m

private def printIdCore (id : Name) : CommandElabM Unit := do
let env ← getEnv
Expand Down
2 changes: 2 additions & 0 deletions src/Lean/PrettyPrinter/Delaborator/FieldNotation.lean
Original file line number Diff line number Diff line change
Expand Up @@ -89,6 +89,8 @@ def fieldNotationCandidate? (f : Expr) (args : Array Expr) (useGeneralizedFieldN
-- Handle structure projections
try
let (field, numParams, _) ← projInfo c
unless numParams + 1 ≤ args.size do return none
unless (← whnf <| ← inferType args[numParams]!).isAppOf c.getPrefix do return none
return (field, numParams)
catch _ => pure ()
-- Handle generalized field notation
Expand Down
62 changes: 62 additions & 0 deletions tests/lean/run/4670.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,62 @@
/-!
# Check types when pretty printing dot notation for structure projections
In type mismatch errors, the 'object of dot notation' might not be valid for dot notation.
-/

structure Foo : Type where
out : Nat

/-!
Was printing `true.out`, but it should have been `Foo.out true`.
-/
/--
error: application type mismatch
Foo.out true
argument
true
has type
Bool : Type
but is expected to have type
Foo : Type
---
info: (sorryAx Foo true).out : Nat
-/
#guard_msgs in #check Foo.out true

/-!
Verifying that generalized field notation does not have this bug.
-/
def Foo.out' (f : Foo) : Nat := f.out
/--
error: application type mismatch
Foo.out' true
argument
true
has type
Bool : Type
but is expected to have type
Foo : Type
---
info: (sorryAx Foo true).out' : Nat
-/
#guard_msgs in #check Foo.out' true

/-!
Verifying that projection notation still pretty prints as normal.
-/
section
variable (f : Foo)
/-- info: f.out : Nat -/
#guard_msgs in #check f.out
end

/-!
Verifying that projection notation still pretty prints through type synonys.
-/
section
def Baz := Foo
variable (f : Baz)
/-- info: f.out : Nat -/
#guard_msgs in #check f.out
end

0 comments on commit 7cd406f

Please sign in to comment.