Skip to content

Commit e0b813b

Browse files
chore: more module system fixes and refinements for finishing batteries port (#10873)
Backport 37b78bd from #10819. Co-authored-by: Sebastian Ullrich <[email protected]>
1 parent 74fea5d commit e0b813b

File tree

20 files changed

+132
-64
lines changed

20 files changed

+132
-64
lines changed

src/Lean/AddDecl.lean

Lines changed: 19 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -96,13 +96,17 @@ def warnIfUsesSorry (decl : Declaration) : CoreM Unit := do
9696
-- This case should not happen, but it ensures a warning will get logged no matter what.
9797
logWarning <| .tagged `hasSorry m!"declaration uses 'sorry'"
9898

99+
builtin_initialize
100+
registerTraceClass `addDecl
101+
99102
/--
100103
Adds the given declaration to the environment's private scope, deriving a suitable presentation in
101104
the public scope if under the module system and if the declaration is not private. If `forceExpose`
102105
is true, exposes the declaration body, i.e. preserves the full representation in the public scope,
103106
independently of `Environment.isExporting` and even for theorems.
104107
-/
105-
def addDecl (decl : Declaration) (forceExpose := false) : CoreM Unit := do
108+
def addDecl (decl : Declaration) (forceExpose := false) : CoreM Unit :=
109+
withTraceNode `addDecl (fun _ => return m!"adding declarations {decl.getNames}") do
106110
-- register namespaces for newly added constants; this used to be done by the kernel itself
107111
-- but that is incompatible with moving it to a separate task
108112
-- NOTE: we do not use `getTopLevelNames` here so that inductive types are registered as
@@ -115,26 +119,37 @@ def addDecl (decl : Declaration) (forceExpose := false) : CoreM Unit := do
115119
let (name, info, kind) ← match decl with
116120
| .thmDecl thm =>
117121
if !forceExpose && (← getEnv).header.isModule then
122+
trace[addDecl] "exporting theorem {thm.name} as axiom"
118123
exportedInfo? := some <| .axiomInfo { thm with isUnsafe := false }
119124
pure (thm.name, .thmInfo thm, .thm)
120125
| .defnDecl defn | .mutualDefnDecl [defn] =>
121126
if !forceExpose && (← getEnv).header.isModule && !(← getEnv).isExporting then
127+
trace[addDecl] "exporting definition {defn.name} as axiom"
122128
exportedInfo? := some <| .axiomInfo { defn with isUnsafe := defn.safety == .unsafe }
123129
pure (defn.name, .defnInfo defn, .defn)
124130
| .opaqueDecl op =>
125131
if !forceExpose && (← getEnv).header.isModule && !(← getEnv).isExporting then
132+
trace[addDecl] "exporting opaque {op.name} as axiom"
126133
exportedInfo? := some <| .axiomInfo { op with }
127134
pure (op.name, .opaqueInfo op, .opaque)
128135
| .axiomDecl ax => pure (ax.name, .axiomInfo ax, .axiom)
129-
| _ => return (← doAdd)
136+
| _ =>
137+
trace[addDecl] "no matching async adding rules, adding synchronously"
138+
return (← doAdd)
130139

131-
if decl.getTopLevelNames.all isPrivateName && !(← ResolveName.backward.privateInPublic.getM) then
132-
exportedInfo? := none
140+
if decl.getTopLevelNames.all isPrivateName then
141+
if (← ResolveName.backward.privateInPublic.getM) then
142+
trace[addDecl] "private decl under `privateInPublic`, exporting as is"
143+
exportedInfo? := some info
144+
else
145+
trace[addDecl] "not exporting private declaration at all"
146+
exportedInfo? := none
133147
else
134148
-- preserve original constant kind in extension if different from exported one
135149
if exportedInfo?.isSome then
136150
modifyEnv (privateConstKindsExt.insert · name kind)
137151
else
152+
trace[addDecl] "no matching exporting rules, exporting as is"
138153
exportedInfo? := some info
139154

140155
-- no environment extension changes to report after kernel checking; ensures we do not

src/Lean/Attributes.lean

Lines changed: 8 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -148,11 +148,14 @@ def throwAttrDeclNotOfExpectedType (attrName declName : Name) (givenType expecte
148148
throwError m!"Cannot add attribute `[{attrName}]`: Declaration `{.ofConstName declName}` has type{indentExpr givenType}\n\
149149
but `[{attrName}]` can only be added to declarations of type{indentExpr expectedType}"
150150

151-
def ensureAttrDeclIsPublic [MonadEnv m] (attrName declName : Name) (attrKind : AttributeKind) : m Unit := do
152-
if (← getEnv).header.isModule && attrKind != .local && !((← getEnv).setExporting true).contains declName then
153-
throwError m!"Cannot add attribute `[{attrName}]`: Declaration `{.ofConstName declName}` must be public"
154-
155-
def ensureAttrDeclIsMeta [MonadEnv m] (attrName declName : Name) (attrKind : AttributeKind) : m Unit := do
151+
def ensureAttrDeclIsPublic (attrName declName : Name) (attrKind : AttributeKind) : AttrM Unit := do
152+
if (← getEnv).header.isModule && attrKind != .local then
153+
withExporting do
154+
checkPrivateInPublic declName
155+
if !(← hasConst declName) then
156+
throwError m!"Cannot add attribute `[{attrName}]`: Declaration `{.ofConstName declName}` must be public"
157+
158+
def ensureAttrDeclIsMeta (attrName declName : Name) (attrKind : AttributeKind) : AttrM Unit := do
156159
if (← getEnv).header.isModule && !isMeta (← getEnv) declName then
157160
throwError m!"Cannot add attribute `[{attrName}]`: Declaration `{.ofConstName declName}` must be marked as `meta`"
158161
-- Make sure attributed decls can't refer to private meta imports, which is already checked for

src/Lean/Compiler/LCNF/Types.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -195,6 +195,10 @@ where
195195
visitApp (f : Expr) (args : Array Expr) := do
196196
let fNew ← match f with
197197
| .const declName us =>
198+
if (← getEnv).isExporting && isPrivateName declName then
199+
-- This branch can happen under `backward.privateInPublic`; restore original behavior of
200+
-- failing here, which is caught and ignored above by `observing`.
201+
throwError "internal compiler error: private in public"
198202
let .inductInfo _ ← getConstInfo declName | return anyExpr
199203
pure <| .const declName us
200204
| .fvar .. => pure f

src/Lean/Elab/App.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1616,7 +1616,7 @@ private def elabAppLValsAux (namedArgs : Array NamedArg) (args : Array Arg) (exp
16161616
| LValResolution.projFn baseStructName structName fieldName =>
16171617
let f ← mkBaseProjections baseStructName structName f
16181618
let some info := getFieldInfo? (← getEnv) baseStructName fieldName | unreachable!
1619-
if isInaccessiblePrivateName (← getEnv) info.projFn then
1619+
if (← isInaccessiblePrivateName info.projFn) then
16201620
throwError "Field `{fieldName}` from structure `{structName}` is private"
16211621
let projFn ← mkConst info.projFn
16221622
let projFn ← addProjTermInfo lval.getRef projFn
@@ -1750,7 +1750,7 @@ where
17501750
match resultTypeFn with
17511751
| .const declName .. =>
17521752
let env ← getEnv
1753-
if isInaccessiblePrivateName env declName then
1753+
if (← isInaccessiblePrivateName declName) then
17541754
throwError "The private declaration `{.ofConstName declName}` is not accessible in the current context"
17551755
-- Recall that the namespace for private declarations is non-private.
17561756
let fullName := privateToUserName declName ++ id

src/Lean/Elab/BuiltinNotation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -56,7 +56,7 @@ open Meta
5656
(fun ival _ => do
5757
match ival.ctors with
5858
| [ctor] =>
59-
if isInaccessiblePrivateName (← getEnv) ctor then
59+
if (← isInaccessiblePrivateName ctor) then
6060
throwError "Invalid `⟨...⟩` notation: Constructor for `{ival.name}` is marked as private"
6161
let cinfo ← getConstInfoCtor ctor
6262
let numExplicitFields ← forallTelescopeReducing cinfo.type fun xs _ => do

src/Lean/Elab/BuiltinTerm.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -164,7 +164,7 @@ private def getMVarFromUserName (ident : Syntax) : MetaM Expr := do
164164
-- `by` switches from an exported to a private context, so we must disallow unassigned
165165
-- metavariables in the goal in this case as they could otherwise leak private data back into
166166
-- the exported context.
167-
mkTacticMVar expectedType stx .term (delayOnMVars := (← getEnv).isExporting)
167+
mkTacticMVar expectedType stx .term (delayOnMVars := (← getEnv).isExporting && !(← backward.proofsInPublic.getM))
168168
| none =>
169169
tryPostpone
170170
throwError ("invalid 'by' tactic, expected type has not been provided")

src/Lean/Elab/MutualDef.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1222,7 +1222,9 @@ where
12221222
assert! view.kind.isTheorem
12231223
let env ← getEnv
12241224
let async ← env.addConstAsync declId.declName .thm
1225-
(exportedKind? := guard (!isPrivateName declId.declName) *> some .axiom)
1225+
(exportedKind? :=
1226+
guard (!isPrivateName declId.declName || (← ResolveName.backward.privateInPublic.getM)) *>
1227+
some .axiom)
12261228
setEnv async.mainEnv
12271229

12281230
-- TODO: parallelize header elaboration as well? Would have to refactor auto implicits catch,

src/Lean/Elab/PreDefinition/Basic.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -135,7 +135,10 @@ private def shouldGenCodeFor (preDef : PreDefinition) : Bool :=
135135
!preDef.kind.isTheorem && !preDef.modifiers.isNoncomputable
136136

137137
private def compileDecl (decl : Declaration) : TermElabM Unit := do
138-
Lean.compileDecl (logErrors := !(← read).isNoncomputableSection) decl
138+
Lean.compileDecl
139+
-- `meta` should disregard `noncomputable section`
140+
(logErrors := !(← read).isNoncomputableSection || decl.getTopLevelNames.any (isMeta (← getEnv)))
141+
decl
139142

140143
register_builtin_option diagnostics.threshold.proofSize : Nat := {
141144
defValue := 16384

src/Lean/Elab/StructInst.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1190,7 +1190,7 @@ private def elabStructInstView (s : StructInstView) (structName : Name) (structT
11901190
TermElabM Expr := withRef s.ref do
11911191
let env ← getEnv
11921192
let ctorVal := getStructureCtor env structName
1193-
if isInaccessiblePrivateName env ctorVal.name then
1193+
if (← isInaccessiblePrivateName ctorVal.name) then
11941194
throwError "invalid \{...} notation, constructor for `{.ofConstName structName}` is marked as private"
11951195
let { ctorFn, ctorFnType, structType, levels, params } ← mkCtorHeader ctorVal structType?
11961196
let (_, fields) ← expandFields structName s.fields (recover := (← read).errToSorry)

src/Lean/Elab/SyntheticMVars.lean

Lines changed: 8 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -340,6 +340,13 @@ private def TacticMVarKind.maybeWithoutRecovery (kind : TacticMVarKind) (m : Tac
340340
else
341341
m
342342

343+
register_builtin_option backward.proofsInPublic : Bool := {
344+
defValue := false
345+
descr := "(module system) Do not abstract proofs used in the public scope into auxiliary \
346+
theorems. Enabling this option may lead to failures or, when `backward.privateInPublic` and \
347+
its `warn` sub-option are enabled, additional warnings from private accesses."
348+
}
349+
343350
mutual
344351

345352
/--
@@ -352,7 +359,7 @@ mutual
352359
partial def runTactic (mvarId : MVarId) (tacticCode : Syntax) (kind : TacticMVarKind) (report := true) : TermElabM Unit := withoutAutoBoundImplicit do
353360
let wasExporting := (← getEnv).isExporting
354361
-- exit exporting context if entering proof
355-
let isNoLongerExporting ← pure wasExporting <&&> do
362+
let isNoLongerExporting ← pure (wasExporting && !(← backward.proofsInPublic.getM)) <&&> do
356363
mvarId.withContext do
357364
isProp (← mvarId.getType)
358365
instantiateMVarDeclMVars mvarId

0 commit comments

Comments
 (0)