Skip to content

Commit

Permalink
feat: hygieneInfo parser (aka this 2.0)
Browse files Browse the repository at this point in the history
  • Loading branch information
digama0 committed May 30, 2023
1 parent e01003f commit 13809d5
Show file tree
Hide file tree
Showing 9 changed files with 50 additions and 5 deletions.
11 changes: 10 additions & 1 deletion src/Init/Meta.lean
Original file line number Diff line number Diff line change
Expand Up @@ -271,10 +271,11 @@ abbrev CharLit := TSyntax charLitKind
abbrev NameLit := TSyntax nameLitKind
abbrev ScientificLit := TSyntax scientificLitKind
abbrev NumLit := TSyntax numLitKind
abbrev HygieneInfo := TSyntax hygieneInfoKind

end Syntax

export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit)
export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit HygieneInfo)

namespace TSyntax

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

def getHygieneInfo (s : HygieneInfo) : Name :=
s.raw[0].getId

namespace Compat

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

end TSyntax

def HygieneInfo.mkIdent (s : HygieneInfo) (val : Name) (canonical := false) : Ident :=
let src := s.raw[0]
let id := { extractMacroScopes src.getId with name := val.eraseMacroScopes }.review
⟨Syntax.ident (SourceInfo.fromRef src canonical) (toString val).toSubstring id []⟩

/-- Reflect a runtime datum back to surface syntax (best-effort). -/
class Quote (α : Type) (k : SyntaxNodeKind := `term) where
quote : α → TSyntax k
Expand Down
10 changes: 10 additions & 0 deletions src/Init/Prelude.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3721,6 +3721,16 @@ abbrev nameLitKind : SyntaxNodeKind := `name
/-- `fieldIdx` is the node kind of projection indices like the `2` in `x.2`. -/
abbrev fieldIdxKind : SyntaxNodeKind := `fieldIdx

/--
`hygieneInfo` is the node kind of the `hygieneInfo` parser, which is an
"invisible token" which captures the hygiene information at the current point
without parsing anything.
They can be used to generate identifiers (with `Lean.HygieneInfo.mkIdent`)
as if they were introduced by the calling context, not the called macro.
-/
abbrev hygieneInfoKind : SyntaxNodeKind := `hygieneInfo

/--
`interpolatedStrLitKind` is the node kind of interpolated string literal
fragments like `"value = {` and `}"` in `s!"value = {x}"`.
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Elab/Syntax.lean
Original file line number Diff line number Diff line change
Expand Up @@ -260,7 +260,7 @@ open Lean.Parser.Command

private def declareSyntaxCatQuotParser (catName : Name) : CommandElabM Unit := do
if let .str _ suffix := catName then
let quotSymbol := "`(" ++ suffix ++ "|"
let quotSymbol := "`(" ++ suffix ++ "| "
let name := catName ++ `quot
let cmd ← `(
@[term_parser] def $(mkIdent name) : Lean.ParserDescr :=
Expand Down
3 changes: 2 additions & 1 deletion src/Lean/Parser.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,7 +27,8 @@ builtin_initialize
register_parser_alias (kind := charLitKind) "char" charLit
register_parser_alias (kind := nameLitKind) "name" nameLit
register_parser_alias (kind := scientificLitKind) "scientific" scientificLit
register_parser_alias (kind := identKind) "ident" ident
register_parser_alias (kind := identKind) ident
register_parser_alias (kind := hygieneInfoKind) hygieneInfo
register_parser_alias "colGt" checkColGt { stackSz? := some 0 }
register_parser_alias "colGe" checkColGe { stackSz? := some 0 }
register_parser_alias "colEq" checkColEq { stackSz? := some 0 }
Expand Down
14 changes: 14 additions & 0 deletions src/Lean/Parser/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1072,6 +1072,20 @@ def identEq (id : Name) : Parser := {
info := mkAtomicInfo "ident"
}

def hygieneInfoFn : ParserFn := fun c s =>
let input := c.input
let pos := s.pos
let str := mkEmptySubstringAt input pos
let info := SourceInfo.original str pos str pos
let ident := mkIdent info str .anonymous
let stx := mkNode hygieneInfoKind #[ident]
s.pushSyntax stx

def hygieneInfoNoAntiquot : Parser := {
fn := hygieneInfoFn
info := nodeInfo hygieneInfoKind epsilonInfo
}

namespace ParserState

def keepTop (s : SyntaxStack) (startStackSize : Nat) : SyntaxStack :=
Expand Down
3 changes: 3 additions & 0 deletions src/Lean/Parser/Extra.lean
Original file line number Diff line number Diff line change
Expand Up @@ -37,6 +37,9 @@ attribute [run_builtin_parser_attribute_hooks]
@[run_builtin_parser_attribute_hooks] def rawIdent : Parser :=
withAntiquot (mkAntiquot "ident" identKind) rawIdentNoAntiquot

@[run_builtin_parser_attribute_hooks] def hygieneInfo : Parser :=
withAntiquot (mkAntiquot "hygieneInfo" hygieneInfoKind) hygieneInfoNoAntiquot

@[run_builtin_parser_attribute_hooks] def numLit : Parser :=
withAntiquot (mkAntiquot "num" numLitKind) numLitNoAntiquot

Expand Down
1 change: 1 addition & 0 deletions src/Lean/PrettyPrinter/Formatter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -489,6 +489,7 @@ def sepByNoAntiquot.formatter (p pSep : Formatter) : Formatter := do
@[combinator_formatter skip] def skip.formatter : Formatter := pure ()

@[combinator_formatter pushNone] def pushNone.formatter : Formatter := goLeft
@[combinator_formatter hygieneInfoNoAntiquot] def hygieneInfoNoAntiquot.formatter : Formatter := goLeft

@[combinator_formatter interpolatedStr]
def interpolatedStr.formatter (p : Formatter) : Formatter := do
Expand Down
1 change: 1 addition & 0 deletions src/Lean/PrettyPrinter/Parenthesizer.lean
Original file line number Diff line number Diff line change
Expand Up @@ -492,6 +492,7 @@ def sepByNoAntiquot.parenthesizer (p pSep : Parenthesizer) : Parenthesizer := do
@[combinator_parenthesizer skip] def skip.parenthesizer : Parenthesizer := pure ()

@[combinator_parenthesizer pushNone] def pushNone.parenthesizer : Parenthesizer := goLeft
@[combinator_parenthesizer hygieneInfoNoAntiquot] def hygieneInfoNoAntiquot.parenthesizer : Parenthesizer := goLeft

@[combinator_parenthesizer interpolatedStr]
def interpolatedStr.parenthesizer (p : Parenthesizer) : Parenthesizer := do
Expand Down
10 changes: 8 additions & 2 deletions tests/lean/interactive/1265.lean.expected.out
Original file line number Diff line number Diff line change
@@ -1,7 +1,10 @@
{"textDocument": {"uri": "file://1265.lean"},
"position": {"line": 0, "character": 51}}
{"items":
[{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
[{"label": "getHygieneInfo",
"kind": 3,
"detail": "Lean.HygieneInfo → Lean.Name"},
{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
{"label": "getName", "kind": 3, "detail": "Lean.NameLit → Lean.Name"},
{"label": "expandInterpolatedStr",
"kind": 3,
Expand All @@ -22,7 +25,10 @@
{"textDocument": {"uri": "file://1265.lean"},
"position": {"line": 2, "character": 53}}
{"items":
[{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
[{"label": "getHygieneInfo",
"kind": 3,
"detail": "Lean.HygieneInfo → Lean.Name"},
{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"},
{"label": "getName", "kind": 3, "detail": "Lean.NameLit → Lean.Name"},
{"label": "expandInterpolatedStr",
"kind": 3,
Expand Down

0 comments on commit 13809d5

Please sign in to comment.