Skip to content

Commit 1bb0f3b

Browse files
committed
feat: hint about inaccessible private declaration on dot notation failure
1 parent 114f7e4 commit 1bb0f3b

File tree

4 files changed

+49
-16
lines changed

4 files changed

+49
-16
lines changed

src/Lean/Elab/App.lean

Lines changed: 14 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -1261,8 +1261,11 @@ inductive LValResolution where
12611261
The `baseName` is the base name of the type to search for in the parameter list. -/
12621262
| localRec (baseName : Name) (fvar : Expr)
12631263

1264+
private def mkLValError (e : Expr) (eType : Expr) (msg : MessageData) : MessageData :=
1265+
m!"{msg}{indentExpr e}\nhas type{indentExpr eType}"
1266+
12641267
private def throwLValErrorAt (ref : Syntax) (e : Expr) (eType : Expr) (msg : MessageData) : TermElabM α :=
1265-
throwErrorAt ref "{msg}{indentExpr e}\nhas type{indentExpr eType}"
1268+
throwErrorAt ref (mkLValError e eType msg)
12661269

12671270
private def throwLValError (e : Expr) (eType : Expr) (msg : MessageData) : TermElabM α := do
12681271
throwLValErrorAt (← getRef) e eType msg
@@ -1329,9 +1332,13 @@ where
13291332
| some (_, p2) => prodArity p2 + 1
13301333

13311334
private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM LValResolution := do
1332-
let throwInvalidFieldAt {α : Type} (ref : Syntax) (fieldName : String) (fullName : Name) : TermElabM α := do
1333-
throwLValErrorAt ref e eType <| ← mkUnknownIdentifierMessage (declHint := fullName)
1334-
m!"Invalid field `{fieldName}`: The environment does not contain `{fullName}`"
1335+
let throwInvalidFieldAt {α : Type} (ref : Syntax) (fieldName : String) (fullName : Name)
1336+
(declHint := Name.anonymous) : TermElabM α := do
1337+
throwErrorAt ref (←
1338+
-- ordering: put decl hint, if any, last
1339+
mkUnknownIdentifierMessageCore (declHint := declHint)
1340+
(mkLValError e eType
1341+
m!"Invalid field `{fieldName}`: The environment does not contain `{fullName}`"))
13351342
if eType.isForall then
13361343
match lval with
13371344
| LVal.fieldName _ fieldName suffix? fullRef =>
@@ -1397,6 +1404,9 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L
13971404
if let some (baseStructName, fullName) ← findMethod? structName (.mkSimple fieldName) then
13981405
return LValResolution.const baseStructName structName fullName
13991406
throwInvalidFieldAt fullRef fieldName fullName
1407+
-- Suggest a potential unreachable private name as hint. This does not cover structure
1408+
-- inheritance, nor `import all`.
1409+
(declHint := (mkPrivateName env structName).mkStr fieldName)
14001410
| none, LVal.fieldName _ _ (some suffix) fullRef =>
14011411
-- This may be a function constant whose implicit arguments have already been filled in:
14021412
let c := e.getAppFn

src/Lean/Elab/Term/TermElabM.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1976,7 +1976,7 @@ where
19761976
let env ← getEnv
19771977
-- check for scope errors before trying auto implicits
19781978
if env.isExporting then
1979-
if let [(npriv, _)] ← withEnv (env.setExporting false) <| resolveGlobalName n then
1979+
if let [(npriv, _)] ← withoutExporting <| resolveGlobalName n then
19801980
throwUnknownIdentifierAt (declHint := npriv) stx m!"Unknown identifier `{.ofConstName n}`"
19811981
if (← read).autoBoundImplicit &&
19821982
!(← read).autoBoundImplicitForbidden n &&

src/Lean/Exception.lean

Lines changed: 17 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -108,17 +108,8 @@ function directly.
108108
protected def «throwNamedErrorAt» [Monad m] [MonadError m] (ref : Syntax) (name : Name) (msg : MessageData) : m α :=
109109
withRef ref <| Lean.throwNamedError name msg
110110

111-
/--
112-
Creates a `MessageData` that is tagged with `unknownIdentifierMessageTag`.
113-
This tag is used by the 'import unknown identifier' code action to detect messages that should
114-
prompt the code action.
115-
The end position of the range of an unknown identifier message should always point at the end of the
116-
unknown identifier.
117-
118-
If `declHint` is specified, a corresponding hint is added to the message in case the name refers to
119-
a private declaration that is not accessible in the current context.
120-
-/
121-
def mkUnknownIdentifierMessage [Monad m] [MonadEnv m] [MonadError m] (msg : MessageData)
111+
/-- Like `mkUnknownIdentifierMessage`, but does not tag the message. -/
112+
def mkUnknownIdentifierMessageCore [Monad m] [MonadEnv m] [MonadError m] (msg : MessageData)
122113
(declHint := Name.anonymous) : m MessageData := do
123114
let mut msg := msg
124115
let env ← getEnv
@@ -134,6 +125,21 @@ def mkUnknownIdentifierMessage [Monad m] [MonadEnv m] [MonadError m] (msg : Mess
134125
msg ++ .note m!"A private declaration `{c}` (from `{mod}`) exists but would need to be public to access here."
135126
else
136127
msg ++ .note m!"A public declaration `{c}` exists but is imported privately; consider adding `public import {mod}`."
128+
return msg
129+
130+
/--
131+
Creates a `MessageData` that is tagged with `unknownIdentifierMessageTag`.
132+
This tag is used by the 'import unknown identifier' code action to detect messages that should
133+
prompt the code action.
134+
The end position of the range of an unknown identifier message should always point at the end of the
135+
unknown identifier.
136+
137+
If `declHint` is specified, a corresponding hint is added to the message in case the name refers to
138+
a private declaration that is not accessible in the current context.
139+
-/
140+
def mkUnknownIdentifierMessage [Monad m] [MonadEnv m] [MonadError m] (msg : MessageData)
141+
(declHint := Name.anonymous) : m MessageData := do
142+
let msg ← mkUnknownIdentifierMessageCore msg declHint
137143
return MessageData.tagged unknownIdentifierMessageTag msg
138144

139145
/--

tests/pkg/module/Module/Basic.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -479,3 +479,20 @@ attribute [simp] f_struct
479479
/-! `initialize` should be run even if imported IR-only. -/
480480
481481
public initialize initialized : Nat ← pure 5
482+
483+
/-! Error message on private dot notation access. -/
484+
485+
public structure S
486+
487+
def S.s := 1
488+
489+
/--
490+
error: Invalid field `s`: The environment does not contain `S.s`
491+
s
492+
has type
493+
S
494+
495+
Note: A private declaration `S.s` (from the current module) exists but would need to be public to access here.
496+
-/
497+
#guard_msgs in
498+
@[expose] public def useS (s : S) := s.s

0 commit comments

Comments
 (0)