@@ -17,36 +17,48 @@ namespace Lean
1717
1818set_option doc.verso true
1919
20- public structure IdentSuggestion where
21- existingToIncorrect : NameMap NameSet := {}
22- incorrectToExisting : NameMap NameSet := {}
23- deriving Inhabited
20+ public abbrev NameMapExtension := PersistentEnvExtension (Name × Array Name) (Name × Array Name) (NameMap NameSet)
2421
25- def IdentSuggestion.add (table : IdentSuggestion) (trueName : Name) (altNames : Array Name) : IdentSuggestion :=
26- { existingToIncorrect :=
27- table.existingToIncorrect.alter trueName fun old =>
28- altNames.foldl (β := NameSet) (init := old.getD {}) fun accum altName =>
29- accum.insert altName
30- incorrectToExisting :=
31- altNames.foldl (init := table.incorrectToExisting) fun accum altName =>
32- accum.alter altName fun old =>
33- old.getD {} |>.insert trueName
34- }
22+ /--
23+ Create the extension mapping from existing identifiers to the incorrect alternatives for which we
24+ want to provide suggestions. This is mostly equivalent to a {name}`MapDeclarationExtension` or the
25+ extensions underlying {name}`ParametricAttribute` attributes, but it differs in allowing
26+ {name}`suggest_for` attributes to be assigned in files other than the ones where they were defined.
27+ -/
28+ def mkExistingToIncorrect : IO NameMapExtension := registerPersistentEnvExtension {
29+ name := `Lean .identifierSuggestForAttr.existingToIncorrect
30+ mkInitial := pure {},
31+ addImportedFn := fun _ => pure {},
32+ addEntryFn := fun table (trueName, altNames) =>
33+ table.alter trueName fun old =>
34+ altNames.foldl (β := NameSet) (init := old.getD {}) fun accum altName =>
35+ accum.insert altName
36+ exportEntriesFn table :=
37+ table.toArray.map (fun (trueName, altNames) =>(trueName, altNames.toArray))
38+ |>.qsort (lt := fun a b => Name.quickLt a.1 b.1 )
39+ }
3540
36- builtin_initialize identifierSuggestionForAttr : PersistentEnvExtension
37- (Name × Array Name) /- Mapping from existing constant to potential incorrect alternative names -/
38- (Name × Array Name) /- Mapping from existing constant to potential incorrect alternative names -/
39- IdentSuggestion /- Two directional mapping between existing identifier and alternative incorrect mappings -/ ←
40- let ext ← registerPersistentEnvExtension {
41- mkInitial := pure {},
42- addImportedFn := fun modules => pure <|
43- (modules.foldl (init := {}) fun accum entries =>
44- entries.foldl (init := accum) fun accum (trueName, altNames) =>
45- accum.add trueName altNames),
46- addEntryFn := fun table (trueName, altNames) => table.add trueName altNames
47- exportEntriesFn table := table.existingToIncorrect.toArray.map fun (trueName, altNames) =>
48- (trueName, altNames.toArray)
49- }
41+ /--
42+ Create the extension mapping incorrect identifiers to the existing identifiers we want to suggest as
43+ replacements. This association is the opposite of the usual mapping for {name}`ParametricAttribute`
44+ attributes.
45+ -/
46+ def mkIncorrectToExisting : IO NameMapExtension := registerPersistentEnvExtension {
47+ name := `Lean .identifierSuggestForAttr.incorrectToExisting
48+ mkInitial := pure {},
49+ addImportedFn := fun _ => pure {},
50+ addEntryFn := fun table (trueName, altNames) =>
51+ altNames.foldl (init := table) fun accum altName =>
52+ accum.alter altName fun old =>
53+ old.getD {} |>.insert trueName
54+ exportEntriesFn table :=
55+ table.toArray.map (fun (altName, trueNames) => (altName, trueNames.toArray))
56+ |>.qsort (lt := fun a b => Name.quickLt a.1 b.1 )
57+ }
58+
59+ builtin_initialize identifierSuggestionsImpl : NameMapExtension × NameMapExtension ←
60+ let existingToIncorrect ← mkExistingToIncorrect
61+ let incorrectToExisting ← mkIncorrectToExisting
5062
5163 registerBuiltinAttribute {
5264 name := `suggest_for ,
@@ -61,36 +73,46 @@ builtin_initialize identifierSuggestionForAttr : PersistentEnvExtension
6173 | _ => throwError " Invalid `[suggest_for]` attribute syntax {repr stx}"
6274 if isPrivateName decl then
6375 throwError " Cannot make suggestions for private names"
64- modifyEnv (ext.addEntry · (decl, altSyntaxIds.map (·.getId)))
76+ let altIds := altSyntaxIds.map (·.getId)
77+ modifyEnv (existingToIncorrect.addEntry · (decl, altIds))
78+ modifyEnv (incorrectToExisting.addEntry · (decl, altIds))
6579 }
6680
67- return ext
81+ return (existingToIncorrect, incorrectToExisting)
6882
6983/--
7084Given a name, find all the stored correct, existing identifiers that mention that name in a
71- {lit }`suggest_for` annotation.
85+ {name }`suggest_for` annotation.
7286-/
7387public def getSuggestions [Monad m] [MonadEnv m] (incorrectName : Name) : m NameSet := do
74- return identifierSuggestionForAttr.getState (← getEnv)
75- |>.incorrectToExisting
76- |>.find? incorrectName
77- |>.getD {}
88+ let env ← getEnv
89+ let { state, importedEntries } := identifierSuggestionsImpl.2 .toEnvExtension.getState env
90+ let localEntries := state.find? incorrectName |>.getD {}
91+ return importedEntries.foldl (init := localEntries) fun results moduleEntry =>
92+ match moduleEntry.binSearch (incorrectName, default) (fun a b => Name.quickLt a.1 b.1 ) with
93+ | none => results
94+ | some (_, extras) => extras.foldl (init := results) fun accum extra => accum.insert extra
7895
7996/--
8097Given a (presumably existing) identifier, find all the {lit}`suggest_for` annotations that were given
8198for that identifier.
8299-/
83100public def getStoredSuggestions [Monad m] [MonadEnv m] (trueName : Name) : m NameSet := do
84- return identifierSuggestionForAttr.getState (← getEnv)
85- |>.existingToIncorrect
86- |>.find? trueName
87- |>.getD {}
101+ let env ← getEnv
102+ let { state, importedEntries } := identifierSuggestionsImpl.1 .toEnvExtension.getState env
103+ let localEntries := state.find? trueName |>.getD {}
104+ return importedEntries.foldl (init := localEntries) fun results moduleEntry =>
105+ match moduleEntry.binSearch (trueName, default) (fun a b => Name.quickLt a.1 b.1 ) with
106+ | none => results
107+ | some (_, extras) => extras.foldl (init := results) fun accum extra => accum.insert extra
88108
89109/--
90110Throw an unknown constant error message, potentially suggesting alternatives using
91111{name}`suggest_for` attributes. (Like {name}`throwUnknownConstantAt` but with suggestions.)
92112
93- The "Unknown constant `<id>`" message will fully qualify the name, whereas the
113+ The replacement will mimic the path structure of the original as much as possible if they share a
114+ path prefix: if there is a suggestion for replacing `Foo.Bar.jazz` with `Foo.Bar.baz`, then
115+ `Bar.jazz` will be replaced by `Bar.baz` unless the resulting constant is ambiguous.
94116-/
95117public def throwUnknownConstantWithSuggestions (constName : Name) (ref? : Option Syntax := .none) : CoreM α := do
96118 let suggestions := (← getSuggestions constName).toArray
0 commit comments