Skip to content

Commit

Permalink
chore: prepare for #1090
Browse files Browse the repository at this point in the history
  • Loading branch information
leodemoura committed Apr 1, 2022
1 parent 8dddb0d commit 68acfc7
Show file tree
Hide file tree
Showing 8 changed files with 73 additions and 10 deletions.
14 changes: 13 additions & 1 deletion src/Init/Meta.lean
Original file line number Diff line number Diff line change
Expand Up @@ -524,7 +524,7 @@ def decodeNatLitVal? (s : String) : Option Nat :=
else if c.isDigit then decodeDecimalLitAux s 0 0
else none

def isLit? (litKind : SyntaxNodeKind) (stx : Syntax) : Option String :=
def isLitCore? (litKind : SyntaxNodeKind) (stx : Syntax) : Option String :=
match stx with
| Syntax.node _ k args =>
if k == litKind && args.size == 1 then
Expand All @@ -535,6 +535,18 @@ def isLit? (litKind : SyntaxNodeKind) (stx : Syntax) : Option String :=
none
| _ => none

def isLit? (litKind : SyntaxNodeKind) (stx : Syntax) : Option String := -- TODO remove staging hack
if litKind == `num || litKind == `numLit then
isLitCore? `num stx <|> isLitCore? `numLit stx
else if litKind == `str || litKind == `strLit then
isLitCore? `str stx <|> isLitCore? `strLit stx
else if litKind == `char || litKind == `charLit then
isLitCore? `char stx <|> isLitCore? `charLit stx
else if litKind == `name || litKind == `nameLit then
isLitCore? `name stx <|> isLitCore? `nameLit stx
else
isLitCore? litKind stx

private def isNatLitAux (litKind : SyntaxNodeKind) (stx : Syntax) : Option Nat :=
match isLit? litKind stx with
| some val => decodeNatLitVal? val
Expand Down
25 changes: 22 additions & 3 deletions src/Init/Prelude.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1873,8 +1873,16 @@ def setKind (stx : Syntax) (k : SyntaxNodeKind) : Syntax :=
| Syntax.node info _ args => Syntax.node info k args
| _ => stx

def isOfKind (stx : Syntax) (k : SyntaxNodeKind) : Bool :=
beq stx.getKind k
def isOfKind (stx : Syntax) (k : SyntaxNodeKind) : Bool := -- TODO remove staging hack
cond (or (beq k `num) (beq k `numLit))
(or (beq stx.getKind `num) (beq stx.getKind `numLit))
(cond (or (beq k `str) (beq k `strLit))
(or (beq stx.getKind `str) (beq stx.getKind `strLit))
(cond (or (beq k `char) (beq k `charLit))
(or (beq stx.getKind `char) (beq stx.getKind `charLit))
(cond (or (beq k `name) (beq k `nameLit))
(or (beq stx.getKind `name) (beq stx.getKind `nameLit))
(beq stx.getKind k))))

def getArg (stx : Syntax) (i : Nat) : Syntax :=
match stx with
Expand Down Expand Up @@ -1916,13 +1924,24 @@ def matchesNull (stx : Syntax) (n : Nat) : Bool :=
def matchesIdent (stx : Syntax) (id : Name) : Bool :=
and stx.isIdent (beq stx.getId id)

def matchesLit (stx : Syntax) (k : SyntaxNodeKind) (val : String) : Bool :=
def matchesLitCore (stx : Syntax) (k : SyntaxNodeKind) (val : String) : Bool :=
match stx with
| Syntax.node _ k' args => and (beq k k') (match args.getD 0 Syntax.missing with
| Syntax.atom _ val' => beq val val'
| _ => false)
| _ => false

def matchesLit (stx : Syntax) (k : SyntaxNodeKind) (val : String) : Bool := -- TODO remove staging hack
cond (or (beq k `num) (beq k `numLit))
(or (matchesLitCore stx `num val) (matchesLitCore stx `numLit val))
(cond (or (beq k `str) (beq k `strLit))
(or (matchesLitCore stx `str val) (matchesLitCore stx `strLit val))
(cond (or (beq k `char) (beq k `charLit))
(or (matchesLitCore stx `char val) (matchesLitCore stx `charLit val))
(cond (or (beq k `name) (beq k `nameLit))
(or (matchesLitCore stx `name val) (matchesLitCore stx `nameLit val))
(matchesLitCore stx k val))))

def setArgs (stx : Syntax) (args : Array Syntax) : Syntax :=
match stx with
| node info k _ => node info k args
Expand Down
10 changes: 10 additions & 0 deletions src/Lean/Elab/BuiltinTerm.lean
Original file line number Diff line number Diff line change
Expand Up @@ -167,6 +167,8 @@ private def mkTacticMVar (type : Expr) (tacticCode : Syntax) : TermElabM Expr :=
| some val => pure $ mkStrLit val
| none => throwIllFormedSyntax

@[builtinTermElab str] def elabStrLit' : TermElab := elabStrLit -- TODO remove staging hack

private def mkFreshTypeMVarFor (expectedType? : Option Expr) : TermElabM Expr := do
let typeMVar ← mkFreshTypeMVar MetavarKind.synthetic
match expectedType? with
Expand All @@ -185,6 +187,8 @@ private def mkFreshTypeMVarFor (expectedType? : Option Expr) : TermElabM Expr :=
registerMVarErrorImplicitArgInfo mvar.mvarId! stx r
return r

@[builtinTermElab num] def elabNumLit' : TermElab := elabNumLit -- TODO remove staging hack

@[builtinTermElab rawNatLit] def elabRawNatLit : TermElab := fun stx expectedType? => do
match stx[1].isNatLit? with
| some val => return mkRawNatLit val
Expand All @@ -202,17 +206,23 @@ def elabScientificLit : TermElab := fun stx expectedType? => do
registerMVarErrorImplicitArgInfo mvar.mvarId! stx r
return r

@[builtinTermElab scientific] def elabScientificLit' : TermElab := elabScientificLit -- TODO remove staging hack

@[builtinTermElab charLit] def elabCharLit : TermElab := fun stx _ => do
match stx.isCharLit? with
| some val => return mkApp (Lean.mkConst ``Char.ofNat) (mkRawNatLit val.toNat)
| none => throwIllFormedSyntax

@[builtinTermElab char] def elabCharLit' : TermElab := elabCharLit -- TODO remove staging hack

/- A literal of type `Name`. -/
@[builtinTermElab quotedName] def elabQuotedName : TermElab := fun stx _ =>
match stx[0].isNameLit? with
| some val => pure $ toExpr val
| none => throwIllFormedSyntax

@[builtinTermElab name] def elabQuotedName' : TermElab := elabQuotedName -- TODO remove staging hack

/--
A resolved name literal. Evaluates to the full name of the given constant if
existent in the current context, or else fails. -/
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Elab/Level.lean
Original file line number Diff line number Diff line change
Expand Up @@ -65,7 +65,7 @@ partial def elabLevel (stx : Syntax) : LevelElabM Level := withRef stx do
return mkLevelIMax' (← elabLevel stx) lvl
else if kind == ``Lean.Parser.Level.hole then
mkFreshLevelMVar
else if kind == numLitKind || kind == `num then -- TODO remove staging hack
else if kind == numLitKind || kind == `num || kind == `numLit then -- TODO remove staging hack
match stx.isNatLit? with
| some val => checkUniverseOffset val; return Level.ofNat val
| none => throwIllFormedSyntax
Expand Down
6 changes: 3 additions & 3 deletions src/Lean/Elab/PatternVar.lean
Original file line number Diff line number Diff line change
Expand Up @@ -209,13 +209,13 @@ partial def collect (stx : Syntax) : M Syntax := withRef stx <| withFreshMacroSc
return stx.setArg 2 lhs |>.setArg 3 rhs
else if k == ``Lean.Parser.Term.inaccessible then
return stx
else if k == strLitKind || k == `str then -- TODO remove staging hack
else if k == strLitKind || k == `str || k == `strLit then -- TODO remove staging hack
return stx
else if k == numLitKind || k == `num then -- TODO remove staging hack
else if k == numLitKind || k == `num || k == `numLit then -- TODO remove staging hack
return stx
else if k == scientificLitKind then
return stx
else if k == charLitKind || k == `char then -- TODO remove staging hack
else if k == charLitKind || k == `char || k == `charLit then -- TODO remove staging hack
return stx
else if k == ``Lean.Parser.Term.quotedName then
/- Quoted names have an elaboration function associated with them, and they will not be macro expanded.
Expand Down
16 changes: 15 additions & 1 deletion src/Lean/Elab/Syntax.lean
Original file line number Diff line number Diff line change
Expand Up @@ -339,8 +339,22 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do
let stx' ← `($[$doc?:docComment]? def $declName:ident : Lean.ParserDescr := ParserDescr.nodeWithAntiquot $(quote (toString declName.getId)) $(quote stxNodeKind) $val)
withMacroExpansion stx stx' <| elabCommand stx'

-- TODO remove staging hack
def normKindCore (k : SyntaxNodeKind) : SyntaxNodeKind :=
if k == `num || k == `numLit then numLitKind
else if k == `char || k == `charLit then charLitKind
else if k == `str || k == `strLit then strLitKind
else if k == `name || k == `nameLit then nameLitKind
else k

-- TODO remove staging hack
def normKind (k : SyntaxNodeKind) : SyntaxNodeKind :=
match k with
| Name.str s "antiquot" .. => normKindCore s ++ `antiquot
| _ => k

def checkRuleKind (given expected : SyntaxNodeKind) : Bool :=
given == expected || given == expected ++ `antiquot
normKind given == normKind expected || normKind given == normKind (expected ++ `antiquot)

def inferMacroRulesAltKind : Syntax → CommandElabM SyntaxNodeKind
| `(matchAltExpr| | $pat:term => $rhs) => do
Expand Down
9 changes: 8 additions & 1 deletion src/Lean/Parser/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1725,7 +1725,14 @@ instance : Coe String Parser := ⟨fun s => symbol s ⟩
produces the syntax tree for `$e`. -/
def mkAntiquot (name : String) (kind : Option SyntaxNodeKind) (anonymous := true) : Parser :=
let kind := (kind.getD Name.anonymous) ++ `antiquot
let nameP := node `antiquotName $ checkNoWsBefore ("no space before ':" ++ name ++ "'") >> symbol ":" >> nonReservedSymbol name
-- TODO remove staging hack
let kindP :=
if name == "strLit" then nonReservedSymbol "strLit" <|> nonReservedSymbol "str"
else if name == "numLit" then nonReservedSymbol "numLit" <|> nonReservedSymbol "num"
else if name == "nameLit" then nonReservedSymbol "nameLit" <|> nonReservedSymbol "name"
else if name == "charLit" then nonReservedSymbol "charLit" <|> nonReservedSymbol "char"
else nonReservedSymbol name
let nameP := node `antiquotName $ checkNoWsBefore ("no space before ':" ++ name ++ "'") >> symbol ":" >> kindP
-- if parsing the kind fails and `anonymous` is true, check that we're not ignoring a different
-- antiquotation kind via `noImmediateColon`
let nameP := if anonymous then nameP <|> checkNoImmediateColon >> pushNone else nameP
Expand Down
1 change: 1 addition & 0 deletions src/Lean/Syntax.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,6 +27,7 @@ def isLitKind (k : SyntaxNodeKind) : Bool :=
k == strLitKind || k == numLitKind || k == charLitKind || k == nameLitKind || k == scientificLitKind
-- TODO remove staging hack
|| k == `str || k == `num || k == `char || k == `name || k == `scientific
|| k == `strLit || k == `numLit || k == `charLit || k == `nameLit || k == `scientificLit

namespace SyntaxNode

Expand Down

0 comments on commit 68acfc7

Please sign in to comment.