Skip to content

Commit 13809d5

Browse files
committed
feat: hygieneInfo parser (aka this 2.0)
1 parent e01003f commit 13809d5

File tree

9 files changed

+50
-5
lines changed

9 files changed

+50
-5
lines changed

src/Init/Meta.lean

+10-1
Original file line numberDiff line numberDiff line change
@@ -271,10 +271,11 @@ abbrev CharLit := TSyntax charLitKind
271271
abbrev NameLit := TSyntax nameLitKind
272272
abbrev ScientificLit := TSyntax scientificLitKind
273273
abbrev NumLit := TSyntax numLitKind
274+
abbrev HygieneInfo := TSyntax hygieneInfoKind
274275

275276
end Syntax
276277

277-
export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit)
278+
export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit HygieneInfo)
278279

279280
namespace TSyntax
280281

@@ -921,6 +922,9 @@ def getChar (s : CharLit) : Char :=
921922
def getName (s : NameLit) : Name :=
922923
s.raw.isNameLit?.getD .anonymous
923924

925+
def getHygieneInfo (s : HygieneInfo) : Name :=
926+
s.raw[0].getId
927+
924928
namespace Compat
925929

926930
scoped instance : CoeTail (Array Syntax) (Syntax.TSepArray k sep) where
@@ -930,6 +934,11 @@ end Compat
930934

931935
end TSyntax
932936

937+
def HygieneInfo.mkIdent (s : HygieneInfo) (val : Name) (canonical := false) : Ident :=
938+
let src := s.raw[0]
939+
let id := { extractMacroScopes src.getId with name := val.eraseMacroScopes }.review
940+
⟨Syntax.ident (SourceInfo.fromRef src canonical) (toString val).toSubstring id []⟩
941+
933942
/-- Reflect a runtime datum back to surface syntax (best-effort). -/
934943
class Quote (α : Type) (k : SyntaxNodeKind := `term) where
935944
quote : α → TSyntax k

src/Init/Prelude.lean

+10
Original file line numberDiff line numberDiff line change
@@ -3721,6 +3721,16 @@ abbrev nameLitKind : SyntaxNodeKind := `name
37213721
/-- `fieldIdx` is the node kind of projection indices like the `2` in `x.2`. -/
37223722
abbrev fieldIdxKind : SyntaxNodeKind := `fieldIdx
37233723

3724+
/--
3725+
`hygieneInfo` is the node kind of the `hygieneInfo` parser, which is an
3726+
"invisible token" which captures the hygiene information at the current point
3727+
without parsing anything.
3728+
3729+
They can be used to generate identifiers (with `Lean.HygieneInfo.mkIdent`)
3730+
as if they were introduced by the calling context, not the called macro.
3731+
-/
3732+
abbrev hygieneInfoKind : SyntaxNodeKind := `hygieneInfo
3733+
37243734
/--
37253735
`interpolatedStrLitKind` is the node kind of interpolated string literal
37263736
fragments like `"value = {` and `}"` in `s!"value = {x}"`.

src/Lean/Elab/Syntax.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -260,7 +260,7 @@ open Lean.Parser.Command
260260

261261
private def declareSyntaxCatQuotParser (catName : Name) : CommandElabM Unit := do
262262
if let .str _ suffix := catName then
263-
let quotSymbol := "`(" ++ suffix ++ "|"
263+
let quotSymbol := "`(" ++ suffix ++ "| "
264264
let name := catName ++ `quot
265265
let cmd ← `(
266266
@[term_parser] def $(mkIdent name) : Lean.ParserDescr :=

src/Lean/Parser.lean

+2-1
Original file line numberDiff line numberDiff line change
@@ -27,7 +27,8 @@ builtin_initialize
2727
register_parser_alias (kind := charLitKind) "char" charLit
2828
register_parser_alias (kind := nameLitKind) "name" nameLit
2929
register_parser_alias (kind := scientificLitKind) "scientific" scientificLit
30-
register_parser_alias (kind := identKind) "ident" ident
30+
register_parser_alias (kind := identKind) ident
31+
register_parser_alias (kind := hygieneInfoKind) hygieneInfo
3132
register_parser_alias "colGt" checkColGt { stackSz? := some 0 }
3233
register_parser_alias "colGe" checkColGe { stackSz? := some 0 }
3334
register_parser_alias "colEq" checkColEq { stackSz? := some 0 }

src/Lean/Parser/Basic.lean

+14
Original file line numberDiff line numberDiff line change
@@ -1072,6 +1072,20 @@ def identEq (id : Name) : Parser := {
10721072
info := mkAtomicInfo "ident"
10731073
}
10741074

1075+
def hygieneInfoFn : ParserFn := fun c s =>
1076+
let input := c.input
1077+
let pos := s.pos
1078+
let str := mkEmptySubstringAt input pos
1079+
let info := SourceInfo.original str pos str pos
1080+
let ident := mkIdent info str .anonymous
1081+
let stx := mkNode hygieneInfoKind #[ident]
1082+
s.pushSyntax stx
1083+
1084+
def hygieneInfoNoAntiquot : Parser := {
1085+
fn := hygieneInfoFn
1086+
info := nodeInfo hygieneInfoKind epsilonInfo
1087+
}
1088+
10751089
namespace ParserState
10761090

10771091
def keepTop (s : SyntaxStack) (startStackSize : Nat) : SyntaxStack :=

src/Lean/Parser/Extra.lean

+3
Original file line numberDiff line numberDiff line change
@@ -37,6 +37,9 @@ attribute [run_builtin_parser_attribute_hooks]
3737
@[run_builtin_parser_attribute_hooks] def rawIdent : Parser :=
3838
withAntiquot (mkAntiquot "ident" identKind) rawIdentNoAntiquot
3939

40+
@[run_builtin_parser_attribute_hooks] def hygieneInfo : Parser :=
41+
withAntiquot (mkAntiquot "hygieneInfo" hygieneInfoKind) hygieneInfoNoAntiquot
42+
4043
@[run_builtin_parser_attribute_hooks] def numLit : Parser :=
4144
withAntiquot (mkAntiquot "num" numLitKind) numLitNoAntiquot
4245

src/Lean/PrettyPrinter/Formatter.lean

+1
Original file line numberDiff line numberDiff line change
@@ -489,6 +489,7 @@ def sepByNoAntiquot.formatter (p pSep : Formatter) : Formatter := do
489489
@[combinator_formatter skip] def skip.formatter : Formatter := pure ()
490490

491491
@[combinator_formatter pushNone] def pushNone.formatter : Formatter := goLeft
492+
@[combinator_formatter hygieneInfoNoAntiquot] def hygieneInfoNoAntiquot.formatter : Formatter := goLeft
492493

493494
@[combinator_formatter interpolatedStr]
494495
def interpolatedStr.formatter (p : Formatter) : Formatter := do

src/Lean/PrettyPrinter/Parenthesizer.lean

+1
Original file line numberDiff line numberDiff line change
@@ -492,6 +492,7 @@ def sepByNoAntiquot.parenthesizer (p pSep : Parenthesizer) : Parenthesizer := do
492492
@[combinator_parenthesizer skip] def skip.parenthesizer : Parenthesizer := pure ()
493493

494494
@[combinator_parenthesizer pushNone] def pushNone.parenthesizer : Parenthesizer := goLeft
495+
@[combinator_parenthesizer hygieneInfoNoAntiquot] def hygieneInfoNoAntiquot.parenthesizer : Parenthesizer := goLeft
495496

496497
@[combinator_parenthesizer interpolatedStr]
497498
def interpolatedStr.parenthesizer (p : Parenthesizer) : Parenthesizer := do

tests/lean/interactive/1265.lean.expected.out

+8-2
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,10 @@
11
{"textDocument": {"uri": "file://1265.lean"},
22
"position": {"line": 0, "character": 51}}
33
{"items":
4-
[{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
4+
[{"label": "getHygieneInfo",
5+
"kind": 3,
6+
"detail": "Lean.HygieneInfo → Lean.Name"},
7+
{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
58
{"label": "getName", "kind": 3, "detail": "Lean.NameLit → Lean.Name"},
69
{"label": "expandInterpolatedStr",
710
"kind": 3,
@@ -22,7 +25,10 @@
2225
{"textDocument": {"uri": "file://1265.lean"},
2326
"position": {"line": 2, "character": 53}}
2427
{"items":
25-
[{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
28+
[{"label": "getHygieneInfo",
29+
"kind": 3,
30+
"detail": "Lean.HygieneInfo → Lean.Name"},
31+
{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
2632
{"label": "getName", "kind": 3, "detail": "Lean.NameLit → Lean.Name"},
2733
{"label": "expandInterpolatedStr",
2834
"kind": 3,

0 commit comments

Comments
 (0)