@@ -27,6 +27,90 @@ Lean does not provide a default premise selector, so this module is intended to
2727with a downstream package which registers a premise selector.
2828-/
2929
30+ namespace Lean.Expr.FoldRelevantConstantsImpl
31+
32+ open Lean Meta
33+
34+ unsafe structure State where
35+ visited : PtrSet Expr := mkPtrSet
36+ visitedConsts : NameHashSet := {}
37+
38+ unsafe abbrev FoldM := StateT State MetaM
39+
40+ unsafe def fold {α : Type } (f : Name → α → MetaM α) (e : Expr) (acc : α) : FoldM α :=
41+ let rec visit (e : Expr) (acc : α) : FoldM α := do
42+ if (← get).visited.contains e then
43+ return acc
44+ modify fun s => { s with visited := s.visited.insert e }
45+ if ← isProof e then
46+ -- Don't visit proofs.
47+ return acc
48+ match e with
49+ | .forallE n d b bi =>
50+ let r ← visit d acc
51+ withLocalDecl n bi d fun x =>
52+ visit (b.instantiate1 x) r
53+ | .lam n d b bi =>
54+ let r ← visit d acc
55+ withLocalDecl n bi d fun x =>
56+ visit (b.instantiate1 x) r
57+ | .mdata _ b => visit b acc
58+ | .letE n t v b nondep =>
59+ let r₁ ← visit t acc
60+ let r₂ ← visit v r₁
61+ withLetDecl n t v (nondep := nondep) fun x =>
62+ visit (b.instantiate1 x) r₂
63+ | .app f a =>
64+ let fi ← getFunInfo f (some 1 )
65+ if fi.paramInfo[0 ]!.isInstImplicit then
66+ -- Don't visit implicit arguments.
67+ visit f acc
68+ else
69+ visit a (← visit f acc)
70+ | .proj _ _ b => visit b acc
71+ | .const c _ =>
72+ if (← get).visitedConsts.contains c then
73+ return acc
74+ else
75+ modify fun s => { s with visitedConsts := s.visitedConsts.insert c }
76+ if ← isInstance c then
77+ return acc
78+ else
79+ f c acc
80+ | _ => return acc
81+ visit e acc
82+
83+ @[inline] unsafe def foldUnsafe {α : Type } (e : Expr) (init : α) (f : Name → α → MetaM α) : MetaM α :=
84+ (fold f e init).run' {}
85+
86+ end FoldRelevantConstantsImpl
87+
88+ /-- Apply `f` to every constant occurring in `e` once, skipping instance arguments and proofs. -/
89+ @[implemented_by FoldRelevantConstantsImpl.foldUnsafe]
90+ public opaque foldRelevantConstants {α : Type } (e : Expr) (init : α) (f : Name → α → MetaM α) : MetaM α := pure init
91+
92+ /-- Collect the constants occuring in `e` (once each), skipping instance arguments and proofs. -/
93+ public def relevantConstants (e : Expr) : MetaM (Array Name) := foldRelevantConstants e #[] (fun n ns => return ns.push n)
94+
95+ /-- Collect the constants occuring in `e` (once each), skipping instance arguments and proofs. -/
96+ public def relevantConstantsAsSet (e : Expr) : MetaM NameSet := foldRelevantConstants e ∅ (fun n ns => return ns.insert n)
97+
98+ end Lean.Expr
99+
100+ open Lean Meta MVarId in
101+ public def Lean.MVarId.getConstants (g : MVarId) : MetaM NameSet := withContext g do
102+ let mut c := (← g.getType).getUsedConstantsAsSet
103+ for t in (← getLocalHyps) do
104+ c := c ∪ (← inferType t).getUsedConstantsAsSet
105+ return c
106+
107+ open Lean Meta MVarId in
108+ public def Lean.MVarId.getRelevantConstants (g : MVarId) : MetaM NameSet := withContext g do
109+ let mut c ← (← g.getType).relevantConstantsAsSet
110+ for t in (← getLocalHyps) do
111+ c := c ∪ (← (← inferType t).relevantConstantsAsSet)
112+ return c
113+
30114@[expose] public section
31115
32116namespace Lean.PremiseSelection
@@ -130,25 +214,37 @@ end Selector
130214
131215section DenyList
132216
133- /-- Premises from a module whose name has one of the following components are not retrieved. -/
217+ /--
218+ Premises from a module whose name has one of the following components are not retrieved.
219+
220+ Use `run_cmd modifyEnv fun env => moduleDenyListExt.addEntry env module` to add a module to the deny list.
221+ -/
134222builtin_initialize moduleDenyListExt : SimplePersistentEnvExtension String (List String) ←
135223 registerSimplePersistentEnvExtension {
136224 addEntryFn := (·.cons)
137- addImportedFn := mkStateFromImportedEntries (·.cons) []
225+ addImportedFn := mkStateFromImportedEntries (·.cons) ["Lake" , "Lean" , "Internal" , "Tactic" ]
138226 }
139227
140- /-- A premise whose name has one of the following components is not retrieved. -/
228+ /--
229+ A premise whose name has one of the following components is not retrieved.
230+
231+ Use `run_cmd modifyEnv fun env => nameDenyListExt.addEntry env name` to add a name to the deny list.
232+ -/
141233builtin_initialize nameDenyListExt : SimplePersistentEnvExtension String (List String) ←
142234 registerSimplePersistentEnvExtension {
143235 addEntryFn := (·.cons)
144- addImportedFn := mkStateFromImportedEntries (·.cons) []
236+ addImportedFn := mkStateFromImportedEntries (·.cons) ["Lake" , "Lean" , "Internal" , "Tactic" ]
145237 }
146238
147- /-- A premise whose `type.getForallBody.getAppFn` is a constant that has one of these prefixes is not retrieved. -/
239+ /--
240+ A premise whose `type.getForallBody.getAppFn` is a constant that has one of these prefixes is not retrieved.
241+
242+ Use `run_cmd modifyEnv fun env => typePrefixDenyListExt.addEntry env typePrefix` to add a type prefix to the deny list.
243+ -/
148244builtin_initialize typePrefixDenyListExt : SimplePersistentEnvExtension Name (List Name) ←
149245 registerSimplePersistentEnvExtension {
150246 addEntryFn := (·.cons)
151- addImportedFn := mkStateFromImportedEntries (·.cons) []
247+ addImportedFn := mkStateFromImportedEntries (·.cons) [`Lake, `Lean ]
152248 }
153249
154250def isDeniedModule (env : Environment) (moduleName : Name) : Bool :=
@@ -157,6 +253,7 @@ def isDeniedModule (env : Environment) (moduleName : Name) : Bool :=
157253def isDeniedPremise (env : Environment) (name : Name) : Bool := Id.run do
158254 if name == ``sorryAx then return true
159255 if name.isInternalDetail then return true
256+ if Lean.Meta.isInstanceCore env name then return true
160257 if (nameDenyListExt.getState env).any (fun p => name.anyS (· == p)) then return true
161258 if let some moduleIdx := env.getModuleIdxFor? name then
162259 let moduleName := env.header.moduleNames[moduleIdx.toNat]!
0 commit comments