Skip to content

Commit

Permalink
chore: mark Meta.Context.config as private (#6051)
Browse files Browse the repository at this point in the history
Motivation: we want to modify the internal representation and improve
`isDefEq` caching.
This PR is preparing the stage for future modifications.
  • Loading branch information
leodemoura authored Nov 13, 2024
1 parent 985600f commit b1e52f1
Show file tree
Hide file tree
Showing 17 changed files with 63 additions and 46 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Elab/PreDefinition/WF/Fix.lean
Original file line number Diff line number Diff line change
Expand Up @@ -227,7 +227,7 @@ def mkFix (preDef : PreDefinition) (prefixArgs : Array Expr) (argsPacker : ArgsP
-- decreasing goals when the function has only one non fixed argument.
-- This renaming is irrelevant if the function has multiple non fixed arguments. See `process*` functions above.
let lctx := (← getLCtx).setUserName x.fvarId! varName
withTheReader Meta.Context (fun ctx => { ctx with lctx }) do
withLCtx' lctx do
let F := xs[1]!
let val := preDef.value.beta (prefixArgs.push x)
let val ← processSumCasesOn x F val fun x F val => do
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Elab/PreDefinition/WF/GuessLex.lean
Original file line number Diff line number Diff line change
Expand Up @@ -166,7 +166,7 @@ def mayOmitSizeOf (is_mutual : Bool) (args : Array Expr) (x : Expr) : MetaM Bool
def withUserNames {α} (xs : Array Expr) (ns : Array Name) (k : MetaM α) : MetaM α := do
let mut lctx ← getLCtx
for x in xs, n in ns do lctx := lctx.setUserName x.fvarId! n
withTheReader Meta.Context (fun ctx => { ctx with lctx }) k
withLCtx' lctx k

/-- Create one measure for each (eligible) parameter of the given predefintion. -/
def simpleMeasures (preDefs : Array PreDefinition) (fixedPrefixSize : Nat)
Expand Down
42 changes: 33 additions & 9 deletions src/Lean/Meta/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -332,7 +332,7 @@ register_builtin_option maxSynthPendingDepth : Nat := {
Contextual information for the `MetaM` monad.
-/
structure Context where
config : Config := {}
private config : Config := {}
/-- Local context -/
lctx : LocalContext := {}
/-- Local instances in `lctx`. -/
Expand Down Expand Up @@ -943,6 +943,15 @@ def elimMVarDeps (xs : Array Expr) (e : Expr) (preserveOrder : Bool := false) :
@[inline] def withConfig (f : Config → Config) : n α → n α :=
mapMetaM <| withReader (fun ctx => { ctx with config := f ctx.config })

@[inline] def withCanUnfoldPred (p : Config → ConstantInfo → CoreM Bool) : n α → n α :=
mapMetaM <| withReader (fun ctx => { ctx with canUnfold? := p })

@[inline] def withIncSynthPending : n α → n α :=
mapMetaM <| withReader (fun ctx => { ctx with synthPendingDepth := ctx.synthPendingDepth + 1 })

@[inline] def withInTypeClassResolution : n α → n α :=
mapMetaM <| withReader (fun ctx => { ctx with inTypeClassResolution := true })

/--
Executes `x` tracking zetaDelta reductions `Config.trackZetaDelta := true`
-/
Expand Down Expand Up @@ -1422,6 +1431,14 @@ def withLocalDecl (name : Name) (bi : BinderInfo) (type : Expr) (k : Expr → n
def withLocalDeclD (name : Name) (type : Expr) (k : Expr → n α) : n α :=
withLocalDecl name BinderInfo.default type k

/--
Similar to `withLocalDecl`, but it does **not** check whether the new variable is a local instance or not.
-/
def withLocalDeclNoLocalInstanceUpdate (name : Name) (bi : BinderInfo) (type : Expr) (x : Expr → MetaM α) : MetaM α := do
let fvarId ← mkFreshFVarId
withReader (fun ctx => { ctx with lctx := ctx.lctx.mkLocalDecl fvarId name type bi }) do
x (mkFVar fvarId)

/-- Append an array of free variables `xs` to the local context and execute `k xs`.
`declInfos` takes the form of an array consisting of:
- the name of the variable
Expand Down Expand Up @@ -1538,11 +1555,11 @@ def withReplaceFVarId {α} (fvarId : FVarId) (e : Expr) : MetaM α → MetaM α
localInstances := ctx.localInstances.erase fvarId }

/--
`withNewMCtxDepth k` executes `k` with a higher metavariable context depth,
where metavariables created outside the `withNewMCtxDepth` (with a lower depth) cannot be assigned.
If `allowLevelAssignments` is set to true, then the level metavariable depth
is not increased, and level metavariables from the outer scope can be
assigned. (This is used by TC synthesis.)
`withNewMCtxDepth k` executes `k` with a higher metavariable context depth,
where metavariables created outside the `withNewMCtxDepth` (with a lower depth) cannot be assigned.
If `allowLevelAssignments` is set to true, then the level metavariable depth
is not increased, and level metavariables from the outer scope can be
assigned. (This is used by TC synthesis.)
-/
def withNewMCtxDepth (k : n α) (allowLevelAssignments := false) : n α :=
mapMetaM (withNewMCtxDepthImp allowLevelAssignments) k
Expand All @@ -1552,13 +1569,20 @@ private def withLocalContextImp (lctx : LocalContext) (localInsts : LocalInstanc
x

/--
`withLCtx lctx localInsts k` replaces the local context and local instances, and then executes `k`.
The local context and instances are restored after executing `k`.
This method assumes that the local instances in `localInsts` are in the local context `lctx`.
`withLCtx lctx localInsts k` replaces the local context and local instances, and then executes `k`.
The local context and instances are restored after executing `k`.
This method assumes that the local instances in `localInsts` are in the local context `lctx`.
-/
def withLCtx (lctx : LocalContext) (localInsts : LocalInstances) : n α → n α :=
mapMetaM <| withLocalContextImp lctx localInsts

/--
Simpler version of `withLCtx` which just updates the local context. It is the resposability of the
caller ensure the local instances are also properly updated.
-/
def withLCtx' (lctx : LocalContext) : n α → n α :=
mapMetaM <| withReader (fun ctx => { ctx with lctx })

/--
Runs `k` in a local environment with the `fvarIds` erased.
-/
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/DiscrTree.lean
Original file line number Diff line number Diff line change
Expand Up @@ -553,8 +553,8 @@ private def getKeyArgs (e : Expr) (isMatch root : Bool) (config : WhnfCoreConfig
if isMatch then
return (.other, #[])
else do
let ctxread
if ctx.config.isDefEqStuckEx then
let cfggetConfig
if cfg.isDefEqStuckEx then
/-
When the configuration flag `isDefEqStuckEx` is set to true,
we want `isDefEq` to throw an exception whenever it tries to assign
Expand Down
16 changes: 8 additions & 8 deletions src/Lean/Meta/ExprDefEq.lean
Original file line number Diff line number Diff line change
Expand Up @@ -364,7 +364,7 @@ private partial def isDefEqBindingAux (lctx : LocalContext) (fvars : Array Expr)
| Expr.forallE n d₁ b₁ _, Expr.forallE _ d₂ b₂ _ => process n d₁ d₂ b₁ b₂
| Expr.lam n d₁ b₁ _, Expr.lam _ d₂ b₂ _ => process n d₁ d₂ b₁ b₂
| _, _ =>
withReader (fun ctx => { ctx with lctx := lctx }) do
withLCtx' lctx do
isDefEqBindingDomain fvars ds₂ do
Meta.isExprDefEqAux (e₁.instantiateRev fvars) (e₂.instantiateRev fvars)

Expand Down Expand Up @@ -758,8 +758,8 @@ mutual
if mvarDecl.depth != (← getMCtx).depth || mvarDecl.kind.isSyntheticOpaque then
traceM `Meta.isDefEq.assign.readOnlyMVarWithBiggerLCtx <| addAssignmentInfo (mkMVar mvarId)
throwCheckAssignmentFailure
let ctxMetareadThe Meta.Context
unless ctxMeta.config.ctxApprox && ctx.mvarDecl.lctx.isSubPrefixOf mvarDecl.lctx do
let cfggetConfig
unless cfg.ctxApprox && ctx.mvarDecl.lctx.isSubPrefixOf mvarDecl.lctx do
traceM `Meta.isDefEq.assign.readOnlyMVarWithBiggerLCtx <| addAssignmentInfo (mkMVar mvarId)
throwCheckAssignmentFailure
/- Create an auxiliary metavariable with a smaller context and "checked" type.
Expand Down Expand Up @@ -814,8 +814,8 @@ mutual

partial def checkApp (e : Expr) : CheckAssignmentM Expr :=
e.withApp fun f args => do
let ctxMetareadThe Meta.Context
if f.isMVar && ctxMeta.config.ctxApprox && args.all Expr.isFVar then
let cfggetConfig
if f.isMVar && cfg.ctxApprox && args.all Expr.isFVar then
let f ← check f
catchInternalId outOfScopeExceptionId
(do
Expand Down Expand Up @@ -1794,8 +1794,8 @@ private partial def isDefEqQuickOther (t s : Expr) : MetaM LBool := do
| LBool.true => return LBool.true
| LBool.false => return LBool.false
| _ =>
let ctxread
if ctx.config.isDefEqStuckEx then do
let cfggetConfig
if cfg.isDefEqStuckEx then do
trace[Meta.isDefEq.stuck] "{t} =?= {s}"
Meta.throwIsDefEqStuck
else
Expand Down Expand Up @@ -1834,7 +1834,7 @@ end
let e ← instantiateMVars e
successK e
else
if (← read).config.isDefEqStuckEx then
if (← getConfig).isDefEqStuckEx then
/-
When `isDefEqStuckEx := true` and `mvar` was created in a previous level,
we should throw an exception. See issue #2736 for a situation where this can happen.
Expand Down
5 changes: 3 additions & 2 deletions src/Lean/Meta/GetUnfoldableConst.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,10 +22,11 @@ private def canUnfoldDefault (cfg : Config) (info : ConstantInfo) : CoreM Bool :

def canUnfold (info : ConstantInfo) : MetaM Bool := do
let ctx ← read
let cfg ← getConfig
if let some f := ctx.canUnfold? then
f ctx.config info
f cfg info
else
canUnfoldDefault ctx.config info
canUnfoldDefault cfg info

/--
Look up a constant name, returning the `ConstantInfo`
Expand Down
7 changes: 1 addition & 6 deletions src/Lean/Meta/InferType.lean
Original file line number Diff line number Diff line change
Expand Up @@ -382,11 +382,6 @@ def isType (e : Expr) : MetaM Bool := do
| .sort .. => return true
| _ => return false

@[inline] private def withLocalDecl' {α} (name : Name) (bi : BinderInfo) (type : Expr) (x : Expr → MetaM α) : MetaM α := do
let fvarId ← mkFreshFVarId
withReader (fun ctx => { ctx with lctx := ctx.lctx.mkLocalDecl fvarId name type bi }) do
x (mkFVar fvarId)

def typeFormerTypeLevelQuick : Expr → Option Level
| .forallE _ _ b _ => typeFormerTypeLevelQuick b
| .sort l => some l
Expand All @@ -403,7 +398,7 @@ where
go (type : Expr) (xs : Array Expr) : MetaM (Option Level) := do
match type with
| .sort l => return some l
| .forallE n d b c => withLocalDecl' n c (d.instantiateRev xs) fun x => go b (xs.push x)
| .forallE n d b c => withLocalDeclNoLocalInstanceUpdate n c (d.instantiateRev xs) fun x => go b (xs.push x)
| _ =>
let type ← whnfD (type.instantiateRev xs)
match type with
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/LazyDiscrTree.lean
Original file line number Diff line number Diff line change
Expand Up @@ -222,8 +222,8 @@ private def getKeyArgs (e : Expr) (isMatch root : Bool) (config : WhnfCoreConfig
if isMatch then
return (.other, #[])
else do
let ctxread
if ctx.config.isDefEqStuckEx then
let cfggetConfig
if cfg.isDefEqStuckEx then
/-
When the configuration flag `isDefEqStuckEx` is set to true,
we want `isDefEq` to throw an exception whenever it tries to assign
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/LevelDefEq.lean
Original file line number Diff line number Diff line change
Expand Up @@ -149,8 +149,8 @@ mutual
if r != LBool.undef then
return r == LBool.true
else if !(← hasAssignableLevelMVar lhs <||> hasAssignableLevelMVar rhs) then
let ctxread
if ctx.config.isDefEqStuckEx && (lhs.isMVar || rhs.isMVar) then do
let cfggetConfig
if cfg.isDefEqStuckEx && (lhs.isMVar || rhs.isMVar) then do
trace[Meta.isLevelDefEq.stuck] "{lhs} =?= {rhs}"
Meta.throwIsDefEqStuck
else
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Match/MatcherApp/Transform.lean
Original file line number Diff line number Diff line change
Expand Up @@ -162,7 +162,7 @@ def refineThrough? (matcherApp : MatcherApp) (e : Expr) :
private def withUserNamesImpl {α} (fvars : Array Expr) (names : Array Name) (k : MetaM α) : MetaM α := do
let lctx := (Array.zip fvars names).foldl (init := ← (getLCtx)) fun lctx (fvar, name) =>
lctx.setUserName fvar.fvarId! name
withTheReader Meta.Context (fun ctx => { ctx with lctx }) k
withLCtx' lctx k

/--
Sets the user name of the FVars in the local context according to the given array of names.
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/SynthInstance.lean
Original file line number Diff line number Diff line change
Expand Up @@ -782,7 +782,7 @@ def synthInstance? (type : Expr) (maxResultSize? : Option Nat := none) : MetaM (
(return m!"{exceptOptionEmoji ·} {← instantiateMVars type}") do
withConfig (fun config => { config with isDefEqStuckEx := true, transparency := TransparencyMode.instances,
foApprox := true, ctxApprox := true, constApprox := false, univApprox := false }) do
withReader (fun ctx => { ctx with inTypeClassResolution := true }) do
withInTypeClassResolution do
let localInsts ← getLocalInstances
let type ← instantiateMVars type
let type ← preprocess type
Expand Down Expand Up @@ -839,7 +839,7 @@ private def synthPendingImp (mvarId : MVarId) : MetaM Bool := withIncRecDepth <|
recordSynthPendingFailure mvarDecl.type
return false
else
withReader (fun ctx => { ctx with synthPendingDepth := ctx.synthPendingDepth + 1 }) do
withIncSynthPending do
trace[Meta.synthPending] "synthPending {mkMVar mvarId}"
let val? ← catchInternalId isDefEqStuckExceptionId (synthInstance? mvarDecl.type (maxResultSize? := none)) (fun _ => pure none)
match val? with
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/Tactic/Intro.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,7 +17,7 @@ namespace Lean.Meta
match i, type with
| 0, type =>
let type := type.instantiateRevRange j fvars.size fvars
withReader (fun ctx => { ctx with lctx := lctx }) do
withLCtx' lctx do
withNewLocalInstances fvars j do
let tag ← mvarId.getTag
let type := type.headBeta
Expand Down Expand Up @@ -57,7 +57,7 @@ namespace Lean.Meta
loop i lctx fvars j s body
else
let type := type.instantiateRevRange j fvars.size fvars
withReader (fun ctx => { ctx with lctx := lctx }) do
withLCtx' lctx do
withNewLocalInstances fvars j do
/- We used to use just `whnf`, but it produces counterintuitive behavior if
- `type` is a metavariable `?m` such that `?m := let x := v; b`, or
Expand Down
3 changes: 0 additions & 3 deletions src/Lean/Meta/Tactic/Rewrites.lean
Original file line number Diff line number Diff line change
Expand Up @@ -60,9 +60,6 @@ private def addImport (name : Name) (constInfo : ConstantInfo) :
pure a
| _ => return #[]

/-- Configuration for `DiscrTree`. -/
def discrTreeConfig : WhnfCoreConfig := {}

/-- Select `=` and `↔` local hypotheses. -/
def localHypotheses (except : List FVarId := []) : MetaM (Array (Expr × Bool × Nat)) := do
let r ← getLocalHyps
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Simp/Rewrite.lean
Original file line number Diff line number Diff line change
Expand Up @@ -552,7 +552,7 @@ private def dischargeUsingAssumption? (e : Expr) : SimpM (Option Expr) := do
partial def dischargeEqnThmHypothesis? (e : Expr) : MetaM (Option Expr) := do
assert! isEqnThmHypothesis e
let mvar ← mkFreshExprSyntheticOpaqueMVar e
withReader (fun ctx => { ctx with canUnfold? := canUnfoldAtMatcher }) do
withCanUnfoldPred canUnfoldAtMatcher do
if let .none ← go? mvar.mvarId! then
instantiateMVars mvar
else
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/UnificationHint.lean
Original file line number Diff line number Diff line change
Expand Up @@ -96,7 +96,7 @@ builtin_initialize

def tryUnificationHints (t s : Expr) : MetaM Bool := do
trace[Meta.isDefEq.hint] "{t} =?= {s}"
unless (← read).config.unificationHints do
unless (← getConfig).unificationHints do
return false
if t.isMVar then
return false
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/WHNF.lean
Original file line number Diff line number Diff line change
Expand Up @@ -529,7 +529,7 @@ private def whnfMatcher (e : Expr) : MetaM Expr := do
TODO: consider other solutions; investigate whether the solution above produces counterintuitive behavior. -/
if (← getTransparency) matches .instances | .reducible then
-- Also unfold some default-reducible constants; see `canUnfoldAtMatcher`
withTransparency .instances <| withReader (fun ctx => { ctx with canUnfold? := canUnfoldAtMatcher }) do
withTransparency .instances <| withCanUnfoldPred canUnfoldAtMatcher do
whnf e
else
-- Do NOT use `canUnfoldAtMatcher` here as it does not affect all/default reducibility and inhibits caching (#2564).
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.lean
Original file line number Diff line number Diff line change
Expand Up @@ -205,7 +205,7 @@ def replaceLPsWithVars (e : Expr) : MetaM Expr := do
| l => if !l.hasParam then some l else none

def isDefEqAssigning (t s : Expr) : MetaM Bool := do
withReader (fun ctx => { ctx with config := { ctx.config with assignSyntheticOpaque := true }}) $
withConfig (fun cfg => { cfg with assignSyntheticOpaque := true }) do
Meta.isDefEq t s

def checkpointDefEq (t s : Expr) : MetaM Bool := do
Expand Down Expand Up @@ -624,7 +624,7 @@ open TopDownAnalyze SubExpr
def topDownAnalyze (e : Expr) : MetaM OptionsPerPos := do
let s₀ ← get
withTraceNode `pp.analyze (fun _ => return e) do
withReader (fun ctx => { ctx with config := Elab.Term.setElabConfig ctx.config }) do
withConfig Elab.Term.setElabConfig do
let ϕ : AnalyzeM OptionsPerPos := do withNewMCtxDepth analyze; pure (← get).annotations
try
let knowsType := getPPAnalyzeKnowsType (← getOptions)
Expand Down

0 comments on commit b1e52f1

Please sign in to comment.