Skip to content

Commit

Permalink
feat: Name.append and macro scopes
Browse files Browse the repository at this point in the history
  • Loading branch information
leodemoura committed Nov 30, 2022
1 parent bc21716 commit a095dab
Show file tree
Hide file tree
Showing 3 changed files with 41 additions and 22 deletions.
48 changes: 33 additions & 15 deletions src/Init/Prelude.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3459,19 +3459,13 @@ instance : BEq Name where
beq := Name.beq

/--
Append two hierarchical names. Example:
```lean
`Lean.Meta ++ `Tactic.simp
```
return `Lean.Meta.Tactic.simp`
This function does not have special support for macro scopes.
See `Name.append`.
-/
protected def append : Name → Name → Name
| n, anonymous => n
| n, str p s => Name.mkStr (Name.append n p) s
| n, num p d => Name.mkNum (Name.append n p) d

instance : Append Name where
append := Name.append
def appendCore : Name → Name → Name
| n, .anonymous => n
| n, .str p s => .str (appendCore n p) s
| n, .num p d => .num (appendCore n p) d

end Name

Expand Down Expand Up @@ -4146,7 +4140,7 @@ def MacroScopesView.review (view : MacroScopesView) : Name :=
match view.scopes with
| List.nil => view.name
| List.cons _ _ =>
let base := (Name.mkStr (hAppend (hAppend (Name.mkStr view.name "_@") view.imported) view.mainModule) "_hyg")
let base := (Name.mkStr (Name.appendCore (Name.appendCore (Name.mkStr view.name "_@") view.imported) view.mainModule) "_hyg")
view.scopes.foldl Name.mkNum base

private def assembleParts : List Name → Name → Name
Expand Down Expand Up @@ -4194,12 +4188,36 @@ def addMacroScope (mainModule : Name) (n : Name) (scp : MacroScope) : Name :=
| true => Name.mkNum n scp
| false =>
{ view with
imported := view.scopes.foldl Name.mkNum (hAppend view.imported view.mainModule)
imported := view.scopes.foldl Name.mkNum (Name.appendCore view.imported view.mainModule)
mainModule := mainModule
scopes := List.cons scp List.nil
}.review
| false =>
Name.mkNum (Name.mkStr (hAppend (Name.mkStr n "_@") mainModule) "_hyg") scp
Name.mkNum (Name.mkStr (Name.appendCore (Name.mkStr n "_@") mainModule) "_hyg") scp

/--
Append two names that may have macro scopes. The macro scopes in `b` are always erased.
If `a` has macro scopes, then the are propagated to result of `append a b`
-/
def Name.append (a b : Name) : Name :=
match a.hasMacroScopes, b.hasMacroScopes with
| true, true =>
/-
TODO: this case should be unreachable. Many interactive tests reach this code. We should fix them.
The following workaround erases `b`s macro scopes and keeps the one from `a`.
-/
let view := extractMacroScopes a
{ view with name := appendCore view.name b.eraseMacroScopes }.review
| true, false =>
let view := extractMacroScopes a
{ view with name := appendCore view.name b }.review
| false, true =>
let view := extractMacroScopes b
{ view with name := appendCore a view.name }.review
| false, false => appendCore a b

instance : Append Name where
append := Name.append

/--
Add a new macro scope onto the name `n`, using the monad state to supply the
Expand Down
13 changes: 7 additions & 6 deletions src/Lean/Elab/Deriving/FromToJson.lean
Original file line number Diff line number Diff line change
Expand Up @@ -48,16 +48,17 @@ def mkToJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do
let header ← mkHeader ``ToJson 1 ctx.typeInfos[0]!
let discrs ← mkDiscrs header indVal
let alts ← mkAlts indVal fun ctor args userNames => do
let ctorStr := ctor.name.eraseMacroScopes.getString!
match args, userNames with
| #[], _ => ``(toJson $(quote ctor.name.getString!))
| #[(x, t)], none => ``(mkObj [($(quote ctor.name.getString!), $(← mkToJson x t))])
| #[], _ => ``(toJson $(quote ctorStr))
| #[(x, t)], none => ``(mkObj [($(quote ctorStr), $(← mkToJson x t))])
| xs, none =>
let xs ← xs.mapM fun (x, t) => mkToJson x t
``(mkObj [($(quote ctor.name.getString!), Json.arr #[$[$xs:term],*])])
``(mkObj [($(quote ctorStr), Json.arr #[$[$xs:term],*])])
| xs, some userNames =>
let xs ← xs.mapIdxM fun idx (x, t) => do
`(($(quote userNames[idx]!.getString!), $(← mkToJson x t)))
``(mkObj [($(quote ctor.name.getString!), mkObj [$[$xs:term],*])])
`(($(quote userNames[idx]!.eraseMacroScopes.getString!), $(← mkToJson x t)))
``(mkObj [($(quote ctorStr), mkObj [$[$xs:term],*])])
let auxTerm ← `(match $[$discrs],* with $alts:matchAlt*)
let auxCmd ←
if ctx.usePartial then
Expand Down Expand Up @@ -177,7 +178,7 @@ where
else
``(none)
let stx ←
`((Json.parseTagged json $(quote ctor.getString!) $(quote ctorInfo.numFields) $(quote userNamesOpt)).bind
`((Json.parseTagged json $(quote ctor.eraseMacroScopes.getString!) $(quote ctorInfo.numFields) $(quote userNamesOpt)).bind
(fun jsons => do
$[let $identNames:ident ← $fromJsons:doExpr]*
return $(mkIdent ctor):ident $identNames*))
Expand Down
2 changes: 1 addition & 1 deletion src/lake
Submodule lake updated 1 files
+4 −2 Lake/Util/Name.lean

0 comments on commit a095dab

Please sign in to comment.