Skip to content

Commit

Permalink
feat: support idents in auto tactics
Browse files Browse the repository at this point in the history
  • Loading branch information
digama0 committed May 2, 2024
1 parent 1d17c7d commit 46bbaab
Show file tree
Hide file tree
Showing 2 changed files with 19 additions and 9 deletions.
22 changes: 13 additions & 9 deletions src/Lean/Elab/Binders.lean
Original file line number Diff line number Diff line change
Expand Up @@ -70,30 +70,34 @@ def kindOfBinderName (binderName : Name) : LocalDeclKind :=
else
.default

partial def quoteAutoTactic : Syntax → TermElabM Syntax
| stx@(.ident ..) => throwErrorAt stx "invalid auto tactic, identifier is not allowed"
partial def quoteAutoTactic : Syntax → CoreM Expr
| .ident _ _ val preresolved =>
return mkApp4 (.const ``Syntax.ident [])
(.const ``SourceInfo.none [])
(.app (.const ``String.toSubstring []) (mkStrLit (toString val)))
(toExpr val)
(toExpr preresolved)
| stx@(.node _ k args) => do
if stx.isAntiquot then
throwErrorAt stx "invalid auto tactic, antiquotation is not allowed"
else
let mut quotedArgs ← `(Array.empty)
let ty := .const ``Syntax []
let mut quotedArgs := mkApp (.const ``Array.empty [.zero]) ty
for arg in args do
if k == nullKind && (arg.isAntiquotSuffixSplice || arg.isAntiquotSplice) then
throwErrorAt arg "invalid auto tactic, antiquotation is not allowed"
else
let quotedArg ← quoteAutoTactic arg
quotedArgs ← `(Array.push $quotedArgs $quotedArg)
`(Syntax.node SourceInfo.none $(quote k) $quotedArgs)
| .atom _ val => `(mkAtom $(quote val))
quotedArgs := mkApp3 (.const ``Array.push [.zero]) ty quotedArgs quotedArg
return mkApp3 (.const ``Syntax.node []) (.const ``SourceInfo.none []) (toExpr k) quotedArgs
| .atom _ val => return .app (.const ``mkAtom []) (toExpr val)
| .missing => throwError "invalid auto tactic, tactic is missing"

def declareTacticSyntax (tactic : Syntax) : TermElabM Name :=
withFreshMacroScope do
let name ← MonadQuotation.addMacroScope `_auto
let type := Lean.mkConst `Lean.Syntax
let tactic ← quoteAutoTactic tactic
let value ← elabTerm tactic type
let value ← instantiateMVars value
let value ← quoteAutoTactic tactic
trace[Elab.autoParam] value
let decl := Declaration.defnDecl { name, levelParams := [], type, value, hints := .opaque,
safety := DefinitionSafety.safe }
Expand Down
6 changes: 6 additions & 0 deletions src/Lean/ToExpr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -172,6 +172,12 @@ instance : ToExpr FVarId where
toTypeExpr := mkConst ``FVarId
toExpr fvarId := mkApp (mkConst ``FVarId.mk) (toExpr fvarId.name)

instance : ToExpr Syntax.Preresolved where
toTypeExpr := .const ``Syntax.Preresolved []
toExpr
| .namespace ns => mkApp (.const ``Syntax.Preresolved.namespace []) (toExpr ns)
| .decl a ls => mkApp2 (.const ``Syntax.Preresolved.decl []) (toExpr a) (toExpr ls)

def Expr.toCtorIfLit : Expr → Expr
| .lit (.natVal v) =>
if v == 0 then mkConst ``Nat.zero
Expand Down

0 comments on commit 46bbaab

Please sign in to comment.