Skip to content

Commit 743d015

Browse files
nomeatakim-em
authored andcommitted
fix: deprecation warning location with field notation (#10826)
This PR fixes the location of the “deprecated constant” and similar error messages on field notation (`e.f`, `(e).f`, `e |>. f`). Fixes #10821.
1 parent f7c4f6a commit 743d015

File tree

5 files changed

+79
-14
lines changed

5 files changed

+79
-14
lines changed

src/Lean/Elab/App.lean

Lines changed: 8 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -1340,7 +1340,7 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L
13401340
throwErrorAt ref m!"{msg}{.nil}"
13411341
if eType.isForall then
13421342
match lval with
1343-
| LVal.fieldName _ fieldName suffix? fullRef =>
1343+
| LVal.fieldName ref fieldName suffix? _fullRef =>
13441344
let fullName := Name.str `Function fieldName
13451345
if (← getEnv).contains fullName then
13461346
return LValResolution.const `Function `Function fullName
@@ -1349,7 +1349,7 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L
13491349
been a field in the `Function` namespace, so we needn't wait to check if this is actually
13501350
a constant. If `suffix?` is non-`none`, we prefer to throw the "unknown constant" error
13511351
(because of monad namespaces like `IO` and auxiliary declarations like `mutual_induct`) -/
1352-
throwInvalidFieldAt fullRef fieldName fullName
1352+
throwInvalidFieldAt ref fieldName fullName
13531353
| .fieldIdx .. =>
13541354
throwLValError e eType "Invalid projection: Projections cannot be used on functions"
13551355
else if eType.getAppFn.isMVar then
@@ -1386,7 +1386,7 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L
13861386
throwError m!"Invalid projection: Index `{idx}` is invalid for this structure; {bounds}"
13871387
++ .note m!"The expression{inlineExpr e}has type{inlineExpr eType}which has only {numFields} {fields}"
13881388
++ tupleHint
1389-
| some structName, LVal.fieldName _ fieldName _ fullRef =>
1389+
| some structName, LVal.fieldName ref fieldName _ _ => withRef ref do
13901390
let env ← getEnv
13911391
if isStructure env structName then
13921392
if let some baseStructName := findField? env structName (Name.mkSimple fieldName) then
@@ -1402,15 +1402,15 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L
14021402
-- Then search the environment
14031403
if let some (baseStructName, fullName) ← findMethod? structName (.mkSimple fieldName) then
14041404
return LValResolution.const baseStructName structName fullName
1405-
throwInvalidFieldAt fullRef fieldName fullName
1405+
throwInvalidFieldAt ref fieldName fullName
14061406
-- Suggest a potential unreachable private name as hint. This does not cover structure
14071407
-- inheritance, nor `import all`.
14081408
(declHint := (mkPrivateName env structName).mkStr fieldName)
1409-
| none, LVal.fieldName _ _ (some suffix) fullRef =>
1409+
| none, LVal.fieldName ref _ (some suffix) _fullRef =>
14101410
-- This may be a function constant whose implicit arguments have already been filled in:
14111411
let c := e.getAppFn
14121412
if c.isConst then
1413-
throwUnknownConstantAt fullRef (c.constName! ++ suffix)
1413+
throwUnknownConstantAt ref (c.constName! ++ suffix)
14141414
else
14151415
throwInvalidFieldNotation e eType
14161416
| _, _ => throwInvalidFieldNotation e eType
@@ -1619,7 +1619,7 @@ private def elabAppLValsAux (namedArgs : Array NamedArg) (args : Array Arg) (exp
16191619
let some info := getFieldInfo? (← getEnv) baseStructName fieldName | unreachable!
16201620
if (← isInaccessiblePrivateName info.projFn) then
16211621
throwError "Field `{fieldName}` from structure `{structName}` is private"
1622-
let projFn ← mkConst info.projFn
1622+
let projFn ← withRef lval.getRef <| mkConst info.projFn
16231623
let projFn ← addProjTermInfo lval.getRef projFn
16241624
if lvals.isEmpty then
16251625
let namedArgs ← addNamedArg namedArgs { name := `self, val := Arg.expr f, suppressDeps := true }
@@ -1629,7 +1629,7 @@ private def elabAppLValsAux (namedArgs : Array NamedArg) (args : Array Arg) (exp
16291629
loop f lvals
16301630
| LValResolution.const baseStructName structName constName =>
16311631
let f ← if baseStructName != structName then mkBaseProjections baseStructName structName f else pure f
1632-
let projFn ← mkConst constName
1632+
let projFn ← withRef lval.getRef <| mkConst constName
16331633
let projFn ← addProjTermInfo lval.getRef projFn
16341634
if lvals.isEmpty then
16351635
let (args, namedArgs) ← addLValArg baseStructName f args namedArgs projFn explicit

tests/lean/1038.lean.expected.out

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1 +1 @@
1-
1038.lean:1:7-1:21: error(lean.unknownIdentifier): Unknown constant `IO.FS.realpath`
1+
1038.lean:1:10-1:12: error(lean.unknownIdentifier): Unknown constant `IO.FS.realpath`

tests/lean/346.lean.expected.out

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,5 @@
1-
346.lean:10:6-10:16: error(lean.unknownIdentifier): Unknown constant `SomeType.b`
2-
346.lean:13:2-13:5: error: Invalid field `z`: The environment does not contain `Nat.z`
1+
346.lean:10:15-10:16: error(lean.unknownIdentifier): Unknown constant `SomeType.b`
2+
346.lean:13:4-13:5: error: Invalid field `z`: The environment does not contain `Nat.z`
33
x
44
has type
55
Nat
Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,5 @@
11
funind_errors.lean:4:7-4:26: error(lean.unknownIdentifier): Unknown identifier `doesNotExist.induct`
2-
funind_errors.lean:22:7-22:23: error(lean.unknownIdentifier): Unknown constant `takeWhile.induct`
2+
funind_errors.lean:22:17-22:23: error(lean.unknownIdentifier): Unknown constant `takeWhile.induct`
33
takeWhile.foo.induct.{u_1} {α : Type u_1} (p : α → Bool) (as : Array α) (motive : Nat → Array α → Prop)
44
(case1 :
55
∀ (i : Nat) (r : Array α) (h : i < as.size),
@@ -10,5 +10,5 @@ takeWhile.foo.induct.{u_1} {α : Type u_1} (p : α → Bool) (as : Array α) (mo
1010
have a := as[i];
1111
¬p a = true → motive i r)
1212
(case3 : ∀ (i : Nat) (r : Array α), ¬i < as.size → motive i r) (i : Nat) (r : Array α) : motive i r
13-
funind_errors.lean:38:7-38:20: error(lean.unknownIdentifier): Unknown constant `idEven.induct`
14-
funind_errors.lean:45:7-45:19: error(lean.unknownIdentifier): Unknown constant `idAcc.induct`
13+
funind_errors.lean:38:14-38:20: error(lean.unknownIdentifier): Unknown constant `idEven.induct`
14+
funind_errors.lean:45:13-45:19: error(lean.unknownIdentifier): Unknown constant `idAcc.induct`

tests/lean/run/issue10821.lean

Lines changed: 65 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,65 @@
1+
set_option linter.deprecated true
2+
3+
/-!
4+
Checks that deprecated names in projection notation cause just
5+
the name to be marked, not the whole expression.
6+
-/
7+
def Nat.inc (n : Nat) : Nat := n + 1
8+
9+
@[deprecated inc (since := "2025-10-17")]
10+
def Nat.inc' (n : Nat) : Nat := n + 1
11+
12+
/--
13+
@ +3:5...9
14+
warning: `Nat.inc'` has been deprecated: Use `Nat.inc` instead
15+
-/
16+
#guard_msgs (positions := true) in
17+
example (n : Nat) : Nat :=
18+
n
19+
|>.inc'
20+
21+
/--
22+
@ +2:7...11
23+
warning: `Nat.inc'` has been deprecated: Use `Nat.inc` instead
24+
-/
25+
#guard_msgs (positions := true) in
26+
example (n : Nat) : Nat :=
27+
n |>.inc'
28+
29+
/--
30+
@ +2:6...10
31+
warning: `Nat.inc'` has been deprecated: Use `Nat.inc` instead
32+
-/
33+
#guard_msgs (positions := true) in
34+
example (n : Nat) : Nat :=
35+
(n).inc'
36+
37+
/--
38+
@ +2:4...8
39+
warning: `Nat.inc'` has been deprecated: Use `Nat.inc` instead
40+
-/
41+
#guard_msgs (positions := true) in
42+
example (n : Nat) : Nat :=
43+
n.inc'
44+
45+
/--
46+
@ +2:4...8
47+
error: Invalid field `incc`: The environment does not contain `Nat.incc`
48+
n
49+
has type
50+
Nat
51+
-/
52+
#guard_msgs (positions := true) in
53+
example (n : Nat) : Nat :=
54+
n.incc
55+
56+
/--
57+
@ +2:4...8
58+
error: Invalid field `incc`: The environment does not contain `Nat.incc`
59+
n
60+
has type
61+
Nat
62+
-/
63+
#guard_msgs (positions := true) in
64+
example (n : Nat) : Nat :=
65+
n.incc.foo

0 commit comments

Comments
 (0)