Skip to content

Commit

Permalink
chore: tidying up Lean.unresolveNameGlobal (#4091)
Browse files Browse the repository at this point in the history
The main loop logic could be simplified, and `if let` could be used to
make control flow more obvious.

Also adds a check for macro scopes to prevent `unresolveNameGlobal` from
returning names with macro scopes in the event there's an alias with
one.

This is a follow up to #3946.
  • Loading branch information
kmill authored May 10, 2024
1 parent e237e12 commit a1be9ec
Show file tree
Hide file tree
Showing 2 changed files with 40 additions and 10 deletions.
28 changes: 18 additions & 10 deletions src/Lean/ResolveName.lean
Original file line number Diff line number Diff line change
Expand Up @@ -384,6 +384,16 @@ After `open Foo open Boo`, we have
def resolveGlobalConstNoOverload [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] (id : Syntax) : m Name := do
ensureNonAmbiguous id (← resolveGlobalConst id)

/--
Finds a name that unambiguously resolves to the given name `n₀`.
Considers suffixes of `n₀` and suffixes of aliases of `n₀` when "unresolving".
Aliases are considered first.
When `fullNames` is true, returns either `n₀` or `_root_.n₀`.
This function is meant to be used for pretty printing.
If `n₀` is an accessible name, then the result will be an accessible name.
-/
def unresolveNameGlobal [Monad m] [MonadResolveName m] [MonadEnv m] (n₀ : Name) (fullNames := false) : m Name := do
if n₀.hasMacroScopes then return n₀
if fullNames then
Expand All @@ -393,21 +403,19 @@ def unresolveNameGlobal [Monad m] [MonadResolveName m] [MonadEnv m] (n₀ : Name
let mut initialNames := (getRevAliases (← getEnv) n₀).toArray
initialNames := initialNames.push (rootNamespace ++ n₀)
for initialName in initialNames do
match (← unresolveNameCore initialName) with
| none => continue
| some n => return n
if let some n ← unresolveNameCore initialName then
return n
return n₀ -- if can't resolve, return the original
where
unresolveNameCore (n : Name) : m (Option Name) := do
if n.hasMacroScopes then return none
let mut revComponents := n.componentsRev
let mut candidate := Name.anonymous
for _ in [:revComponents.length] do
match revComponents with
| [] => return none
| cmpt::rest => candidate := Name.appendCore cmpt candidate; revComponents := rest
match (← resolveGlobalName candidate) with
| [(potentialMatch, _)] => if potentialMatch == n₀ then return some candidate else continue
| _ => continue
for cmpt in revComponents do
candidate := Name.appendCore cmpt candidate
if let [(potentialMatch, _)] ← resolveGlobalName candidate then
if potentialMatch == n₀ then
return some candidate
return none

end Lean
22 changes: 22 additions & 0 deletions tests/lean/run/2291.lean
Original file line number Diff line number Diff line change
@@ -1,3 +1,5 @@
import Lean.Elab.Command
import Lean.Elab.Open
/-!
Issue #2291
Expand All @@ -6,3 +8,23 @@ The following example would cause the pretty printer to panic.

set_option trace.compiler.simp true in
#eval [0]


/-!
Fixing the above involved changing `Lean.unresolveNameGlobal`.
Here, we also verify that we do not pretty print using any aliases that have macro scopes.
-/

open Lean in
elab "add_bad_alias " n:ident : command => withFreshMacroScope do
let declName ← Elab.OpenDecl.resolveNameUsingNamespaces [← getCurrNamespace] n
let badName ← MonadQuotation.addMacroScope `bad
modify fun s => { s with env := addAlias s.env badName declName }

def f := 1

add_bad_alias f

-- Formerly was info: bad✝ : ℕ
/-- info: f : Nat -/
#guard_msgs in #check (f)

0 comments on commit a1be9ec

Please sign in to comment.