Skip to content

Commit 4a83264

Browse files
committed
feat: backward.privateInPublic option
1 parent 803ec8f commit 4a83264

File tree

9 files changed

+114
-52
lines changed

9 files changed

+114
-52
lines changed

src/Lean/AddDecl.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -130,7 +130,7 @@ def addDecl (decl : Declaration) (forceExpose := false) : CoreM Unit := do
130130
| .axiomDecl ax => pure (ax.name, .axiomInfo ax, .axiom)
131131
| _ => return (← doAdd)
132132

133-
if decl.getTopLevelNames.all isPrivateName then
133+
if decl.getTopLevelNames.all isPrivateName && !(← ResolveName.backward.privateInPublic.getM) then
134134
exportedInfo? := none
135135
else
136136
-- preserve original constant kind in extension if different from exported one

src/Lean/Data/Options.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -121,6 +121,9 @@ protected def get? [KVMap.Value α] (opts : Options) (opt : Lean.Option α) : Op
121121
protected def get [KVMap.Value α] (opts : Options) (opt : Lean.Option α) : α :=
122122
opts.get opt.name opt.defValue
123123

124+
protected def getM [Monad m] [MonadOptions m] [KVMap.Value α] (opt : Lean.Option α) : m α :=
125+
return opt.get (← getOptions)
126+
124127
protected def set [KVMap.Value α] (opts : Options) (opt : Lean.Option α) (val : α) : Options :=
125128
opts.set opt.name val
126129

src/Lean/Elab/App.lean

Lines changed: 4 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1276,7 +1276,9 @@ private partial def findMethod? (structName fieldName : Name) : MetaM (Option (N
12761276
let find? structName' : MetaM (Option (Name × Name)) := do
12771277
let fullName := privateToUserName structName' ++ fieldName
12781278
-- We do not want to make use of the current namespace for resolution.
1279-
let candidates := ResolveName.resolveGlobalName (← getEnv) Name.anonymous (← getOpenDecls) fullName
1279+
let candidates :=
1280+
(← withTheReader Core.Context ({ · with currNamespace := .anonymous }) do
1281+
resolveGlobalName fullName)
12801282
|>.filter (fun (_, fieldList) => fieldList.isEmpty)
12811283
|>.map Prod.fst
12821284
match candidates with
@@ -1746,7 +1748,7 @@ where
17461748
-- Recall that the namespace for private declarations is non-private.
17471749
let fullName := privateToUserName declName ++ id
17481750
-- Resolve the name without making use of the current namespace, like in `findMethod?`.
1749-
let candidates := ResolveName.resolveGlobalName env Name.anonymous (← getOpenDecls) fullName
1751+
let candidates := ResolveName.resolveGlobalName env (← getOptions) Name.anonymous (← getOpenDecls) fullName
17501752
|>.filter (fun (_, fieldList) => fieldList.isEmpty)
17511753
|>.map Prod.fst
17521754
if !candidates.isEmpty then

src/Lean/Elab/Open.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,7 @@ instance : MonadResolveName (M (m := m)) where
3434
getCurrNamespace := return (← get).currNamespace
3535
getOpenDecls := return (← get).openDecls
3636

37-
def resolveId [MonadResolveName m] (ns : Name) (idStx : Syntax) : m Name := do
37+
def resolveId [MonadOptions m] [MonadResolveName m] (ns : Name) (idStx : Syntax) : m Name := do
3838
let declName := ns ++ idStx.getId
3939
if (← getEnv).contains declName then
4040
return declName
@@ -49,7 +49,7 @@ Uniquely resolves the identifier `idStx` in the provided namespaces `nss`.
4949
5050
If the identifier does not indicate a name in exactly one of the namespaces, an exception is thrown.
5151
-/
52-
def resolveNameUsingNamespacesCore [MonadResolveName m]
52+
def resolveNameUsingNamespacesCore [MonadOptions m] [MonadResolveName m]
5353
(nss : List Name) (idStx : Syntax) : m Name := do
5454
let mut exs := #[]
5555
let mut result := #[]
@@ -70,7 +70,7 @@ def resolveNameUsingNamespacesCore [MonadResolveName m]
7070
else
7171
withRef idStx do throwError "ambiguous identifier `{idStx.getId}`, possible interpretations: {result.map mkConst}"
7272

73-
def elabOpenDecl [MonadResolveName m] [MonadInfoTree m] (stx : TSyntax ``Parser.Command.openDecl) : m (List OpenDecl) := do
73+
def elabOpenDecl [MonadOptions m] [MonadResolveName m] [MonadInfoTree m] (stx : TSyntax ``Parser.Command.openDecl) : m (List OpenDecl) := do
7474
StateRefT'.run' (s := { openDecls := (← getOpenDecls), currNamespace := (← getCurrNamespace) }) do
7575
match stx with
7676
| `(Parser.Command.openDecl| $nss*) =>
@@ -109,7 +109,7 @@ def elabOpenDecl [MonadResolveName m] [MonadInfoTree m] (stx : TSyntax ``Parser.
109109
| _ => throwUnsupportedSyntax
110110
return (← get).openDecls
111111

112-
def resolveNameUsingNamespaces [MonadResolveName m] (nss : List Name) (idStx : Ident) : m Name := do
112+
def resolveNameUsingNamespaces [MonadOptions m] [MonadResolveName m] (nss : List Name) (idStx : Ident) : m Name := do
113113
StateRefT'.run' (s := { openDecls := (← getOpenDecls), currNamespace := (← getCurrNamespace) }) do
114114
resolveNameUsingNamespacesCore (m := M) nss idStx
115115

src/Lean/Elab/Term/TermElabM.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1974,7 +1974,7 @@ where
19741974
let env ← getEnv
19751975
-- check for scope errors before trying auto implicits
19761976
if env.isExporting then
1977-
if let [(npriv, _)] ← withEnv (env.setExporting false) <| resolveGlobalName n then
1977+
if let [(npriv, _)] ← withEnv (env.setExporting false) <| resolveGlobalName (enableLog := false) n then
19781978
throwUnknownIdentifierAt (declHint := npriv) stx m!"Unknown identifier `{.ofConstName n}`"
19791979
if (← read).autoBoundImplicit &&
19801980
!(← read).autoBoundImplicitForbidden n &&

src/Lean/Elab/Util.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -179,6 +179,7 @@ instance (m n) [MonadLift m n] [MonadQuotation n] [MonadMacroAdapter m] : MonadM
179179

180180
def liftMacroM [Monad m] [MonadMacroAdapter m] [MonadEnv m] [MonadRecDepth m] [MonadError m] [MonadResolveName m] [MonadTrace m] [MonadOptions m] [AddMessageContext m] [MonadLiftT IO m] (x : MacroM α) : m α := do
181181
let env ← getEnv
182+
let opts ← getOptions
182183
let currNamespace ← getCurrNamespace
183184
let openDecls ← getOpenDecls
184185
let methods := Macro.mkMethods {
@@ -190,7 +191,7 @@ def liftMacroM [Monad m] [MonadMacroAdapter m] [MonadEnv m] [MonadRecDepth m] [M
190191
hasDecl := fun declName => return env.contains declName
191192
getCurrNamespace := return currNamespace
192193
resolveNamespace := fun n => return ResolveName.resolveNamespace env currNamespace openDecls n
193-
resolveGlobalName := fun n => return ResolveName.resolveGlobalName env currNamespace openDecls n
194+
resolveGlobalName := fun n => return ResolveName.resolveGlobalName env opts currNamespace openDecls n
194195
}
195196
match x { methods := methods
196197
ref := ← getRef

src/Lean/Parser/Extension.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -664,7 +664,7 @@ inductive ParserResolution where
664664
| alias (p : ParserAliasValue)
665665

666666
/-- Resolve the given parser name and return a list of candidates. -/
667-
def resolveParserNameCore (env : Environment) (currNamespace : Name)
667+
private def resolveParserNameCore (env : Environment) (opts : Options) (currNamespace : Name)
668668
(openDecls : List OpenDecl) (ident : Ident) : List ParserResolution := Id.run do
669669
let ⟨.ident (val := val) (preresolved := pre) ..⟩ := ident | return []
670670

@@ -685,7 +685,7 @@ def resolveParserNameCore (env : Environment) (currNamespace : Name)
685685
if isParserCategory env erased then
686686
return [.category erased]
687687

688-
let resolved ← ResolveName.resolveGlobalName env currNamespace openDecls val |>.filterMap fun
688+
let resolved ← ResolveName.resolveGlobalName env opts currNamespace openDecls val |>.filterMap fun
689689
| (name, []) => (isParser name).map fun isDescr => .parser name isDescr
690690
| _ => none
691691
unless resolved.isEmpty do
@@ -699,11 +699,11 @@ def resolveParserNameCore (env : Environment) (currNamespace : Name)
699699

700700
/-- Resolve the given parser name and return a list of candidates. -/
701701
def ParserContext.resolveParserName (ctx : ParserContext) (id : Ident) : List ParserResolution :=
702-
Parser.resolveParserNameCore ctx.env ctx.currNamespace ctx.openDecls id
702+
Parser.resolveParserNameCore ctx.env ctx.options ctx.currNamespace ctx.openDecls id
703703

704704
/-- Resolve the given parser name and return a list of candidates. -/
705705
def resolveParserName (id : Ident) : CoreM (List ParserResolution) :=
706-
return resolveParserNameCore (← getEnv) (← getCurrNamespace) (← getOpenDecls) id
706+
return resolveParserNameCore (← getEnv) (← getOptions) (← getCurrNamespace) (← getOpenDecls) id
707707

708708
def parserOfStackFn (offset : Nat) : ParserFn := fun ctx s => Id.run do
709709
let stack := s.stxStack

0 commit comments

Comments
 (0)