Skip to content

Commit 7ab9ec0

Browse files
committed
feat: flag on TermInfo to force rendering of term in hover
1 parent 10d6232 commit 7ab9ec0

File tree

4 files changed

+31
-20
lines changed

4 files changed

+31
-20
lines changed

src/Lean/Elab/App.lean

Lines changed: 8 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -1581,14 +1581,15 @@ where
15811581

15821582
/-- Adds the `TermInfo` for the field of a projection. See `Lean.Parser.Term.identProjKind`. -/
15831583
private def addProjTermInfo
1584-
(stx : Syntax)
1585-
(e : Expr)
1586-
(expectedType? : Option Expr := none)
1587-
(lctx? : Option LocalContext := none)
1588-
(elaborator : Name := Name.anonymous)
1589-
(isBinder force : Bool := false)
1584+
(stx : Syntax)
1585+
(e : Expr)
1586+
(expectedType? : Option Expr := none)
1587+
(lctx? : Option LocalContext := none)
1588+
(elaborator : Name := Name.anonymous)
1589+
(isBinder force : Bool := false)
1590+
(isDisplayableTerm : Bool := false)
15901591
: TermElabM Expr :=
1591-
addTermInfo (Syntax.node .none Parser.Term.identProjKind #[stx]) e expectedType? lctx? elaborator isBinder force
1592+
addTermInfo (Syntax.node .none Parser.Term.identProjKind #[stx]) e expectedType? lctx? elaborator isBinder force isDisplayableTerm
15921593

15931594
private def elabAppLValsAux (namedArgs : Array NamedArg) (args : Array Arg) (expectedType? : Option Expr) (explicit ellipsis : Bool)
15941595
(f : Expr) (lvals : List LVal) : TermElabM Expr :=

src/Lean/Elab/InfoTree/Types.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -74,6 +74,8 @@ structure TermInfo extends ElabInfo where
7474
expectedType? : Option Expr
7575
expr : Expr
7676
isBinder : Bool := false
77+
/-- Whether `expr` should always be displayed in the language server, e.g. in hovers. -/
78+
isDisplayableTerm : Bool := false
7779
deriving Inhabited
7880

7981
/--

src/Lean/Elab/Term/TermElabM.lean

Lines changed: 11 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -1320,13 +1320,13 @@ def isTacticOrPostponedHole? (e : Expr) : TermElabM (Option MVarId) := do
13201320
| _ => pure none
13211321

13221322
def mkTermInfo (elaborator : Name) (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none)
1323-
(lctx? : Option LocalContext := none) (isBinder := false) :
1323+
(lctx? : Option LocalContext := none) (isBinder := false) (isDisplayableTerm := false) :
13241324
TermElabM (Sum Info MVarId) := do
13251325
match (← isTacticOrPostponedHole? e) with
13261326
| some mvarId => return Sum.inr mvarId
13271327
| none =>
13281328
let e := removeSaveInfoAnnotation e
1329-
return Sum.inl <| Info.ofTermInfo { elaborator, lctx := lctx?.getD (← getLCtx), expr := e, stx, expectedType?, isBinder }
1329+
return Sum.inl <| Info.ofTermInfo { elaborator, lctx := lctx?.getD (← getLCtx), expr := e, stx, expectedType?, isBinder, isDisplayableTerm }
13301330

13311331
def mkPartialTermInfo (elaborator : Name) (stx : Syntax) (expectedType? : Option Expr := none)
13321332
(lctx? : Option LocalContext := none) :
@@ -1350,18 +1350,21 @@ is a constant they will see the constant's doc string.
13501350
-/
13511351
def addTermInfo (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none)
13521352
(lctx? : Option LocalContext := none) (elaborator := Name.anonymous)
1353-
(isBinder := false) (force := false) : TermElabM Expr := do
1353+
(isBinder := false) (force := false) (isDisplayableTerm := false): TermElabM Expr := do
13541354
if (← read).inPattern && !force then
13551355
return mkPatternWithRef e stx
13561356
else
13571357
discard <| withInfoContext'
13581358
(pure ())
1359-
(fun _ => mkTermInfo elaborator stx e expectedType? lctx? isBinder)
1359+
(fun _ => mkTermInfo elaborator stx e expectedType? lctx? isBinder isDisplayableTerm)
13601360
(mkPartialTermInfo elaborator stx expectedType? lctx?)
13611361
return e
13621362

1363-
def addTermInfo' (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none) (lctx? : Option LocalContext := none) (elaborator := Name.anonymous) (isBinder := false) : TermElabM Unit :=
1363+
def addTermInfo' (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none)
1364+
(lctx? : Option LocalContext := none) (elaborator := Name.anonymous) (isBinder := false)
1365+
(isDisplayableTerm := false) : TermElabM Unit :=
13641366
discard <| addTermInfo stx e expectedType? lctx? elaborator isBinder
1367+
(isDisplayableTerm := isDisplayableTerm)
13651368

13661369
def withInfoContext' (stx : Syntax) (x : TermElabM Expr)
13671370
(mkInfo : Expr → TermElabM (Sum Info MVarId)) (mkInfoOnError : TermElabM Info) :
@@ -1389,10 +1392,11 @@ def getBodyInfo? : Info → Option BodyInfo
13891392

13901393
def withTermInfoContext' (elaborator : Name) (stx : Syntax) (x : TermElabM Expr)
13911394
(expectedType? : Option Expr := none) (lctx? : Option LocalContext := none)
1392-
(isBinder : Bool := false) :
1395+
(isBinder : Bool := false) (isDisplayableTerm : Bool := false):
13931396
TermElabM Expr :=
13941397
withInfoContext' stx x
1395-
(mkTermInfo elaborator stx (expectedType? := expectedType?) (lctx? := lctx?) (isBinder := isBinder))
1398+
(mkTermInfo elaborator stx (expectedType? := expectedType?) (lctx? := lctx?)
1399+
(isBinder := isBinder) (isDisplayableTerm := isDisplayableTerm))
13961400
(mkPartialTermInfo elaborator stx (expectedType? := expectedType?) (lctx? := lctx?))
13971401

13981402
/--

src/Lean/Server/InfoUtils.lean

Lines changed: 10 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -393,15 +393,19 @@ where
393393
let eFmt ← PrettyPrinter.ppSignature c
394394
return (some { eFmt with fmt := f!"```lean\n{eFmt.fmt}\n```" }, ← fmtModule? c)
395395
let eFmt ← Meta.ppExpr e
396+
let lctx ← getLCtx
396397
-- Try not to show too scary internals
397398
let showTerm :=
398-
if let .fvar _ := e then
399-
if let some ldecl := (← getLCtx).findFVar? e then
400-
!ldecl.userName.hasMacroScopes
401-
else
402-
false
399+
if ti.isDisplayableTerm then
400+
true
403401
else
404-
isAtomicFormat eFmt
402+
if let .fvar _ := e then
403+
if let some ldecl := lctx.findFVar? e then
404+
!ldecl.userName.hasMacroScopes
405+
else
406+
false
407+
else
408+
isAtomicFormat eFmt
405409
let fmt := if showTerm then f!"{eFmt} : {tpFmt}" else tpFmt
406410
return (some f!"```lean\n{fmt}\n```", none)
407411
| Info.ofFieldInfo fi =>

0 commit comments

Comments
 (0)