Skip to content

Commit 065094f

Browse files
committed
feat: set_config for setting grind configuration options
This PR adds the `set_config` tactic for setting `grind` configuration options. It uses the same syntax used for setting configuration options in the `grind` main tactic.
1 parent 808d123 commit 065094f

File tree

4 files changed

+38
-24
lines changed

4 files changed

+38
-24
lines changed

src/Init/Grind/Interactive.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -240,6 +240,11 @@ syntax (name := exposeNames) "expose_names" : grind
240240
but it sets the option only within the tactics `tacs`. -/
241241
syntax (name := setOption) "set_option " (ident (noWs "." noWs ident)?) ppSpace optionValue " in " grindSeq : grind
242242

243+
/--
244+
`set_config configItem+ in tacs` executes `tacs` with the updated configuration options `configItem+`
245+
-/
246+
syntax (name := setConfig) "set_config " configItem+ " in " grindSeq : grind
247+
243248
/--
244249
Proves `<term>` using the current `grind` state and default search strategy.
245250
-/

src/Lean/Elab/Tactic/Grind/BuiltinTactic.lean

Lines changed: 9 additions & 23 deletions
Original file line numberDiff line numberDiff line change
@@ -426,30 +426,16 @@ where
426426
liftGrindM <| resetAnchors
427427
replaceMainGoal [{ goal with mvarId }]
428428

429-
def isGrindConfigField? (stx : Syntax) : CoreM (Option Name) := do
430-
unless stx.isIdent do return none
431-
let id := stx.getId
432-
let env ← getEnv
433-
let info := getStructureInfo env ``Grind.Config
434-
unless info.fieldNames.contains id do return none
435-
return some id
436-
437429
@[builtin_grind_tactic setOption] def elabSetOption : GrindTactic := fun stx => do
438-
if let some fieldName ← isGrindConfigField? stx[1] then
439-
let val := stx[3]
440-
let val ← match val.isNatLit? with
441-
| some num => pure <| DataValue.ofNat num
442-
| none => match val with
443-
| Syntax.atom _ "true" => pure <| DataValue.ofBool true
444-
| Syntax.atom _ "false" => pure <| DataValue.ofBool false
445-
| _ => throwErrorAt val "`grind` configuration option must be a Boolean or a numeral"
446-
let config := (← read).ctx.config
447-
let config ← setConfigField config fieldName val
448-
withReader (fun c => { c with ctx.config := config }) do
449-
evalGrindTactic stx[5]
450-
else
451-
let options ← Elab.elabSetOption stx[1] stx[3]
452-
withOptions (fun _ => options) do evalGrindTactic stx[5]
430+
let options ← Elab.elabSetOption stx[1] stx[3]
431+
withOptions (fun _ => options) do evalGrindTactic stx[5]
432+
433+
@[builtin_grind_tactic setConfig] def elabSetConfig : GrindTactic := fun stx => do
434+
let `(grind| set_config $[$items:configItem]* in $seq:grindSeq) := stx | throwUnsupportedSyntax
435+
let config := (← read).ctx.config
436+
let config ← elabConfigItems config items
437+
withReader (fun c => { c with ctx.config := config }) do
438+
evalGrindTactic seq
453439

454440
@[builtin_grind_tactic mbtc] def elabMBTC : GrindTactic := fun _ => do
455441
liftGoalM do

src/Lean/Elab/Tactic/Grind/Config.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -13,4 +13,20 @@ namespace Lean.Elab.Tactic.Grind
1313
/-- Sets a field of the `grind` configuration object. -/
1414
declare_config_getter setConfigField Grind.Config
1515

16+
def elabConfigItems (init : Grind.Config) (items : Array (TSyntax `Lean.Parser.Tactic.configItem))
17+
: CoreM Grind.Config := do
18+
let mut config := init
19+
for item in items do
20+
match item with
21+
| `(Lean.Parser.Tactic.configItem| ($fieldName:ident := true))
22+
| `(Lean.Parser.Tactic.configItem| +$fieldName:ident) =>
23+
config ← setConfigField config fieldName.getId true
24+
| `(Lean.Parser.Tactic.configItem| ($fieldName:ident := false))
25+
| `(Lean.Parser.Tactic.configItem| -$fieldName:ident) =>
26+
config ← setConfigField config fieldName.getId false
27+
| `(Lean.Parser.Tactic.configItem| ($fieldName:ident := $val:num)) =>
28+
config ← setConfigField config fieldName.getId (.ofNat val.getNat)
29+
| _ => throwErrorAt item "unexpected configuration option"
30+
return config
31+
1632
end Lean.Elab.Tactic.Grind

tests/lean/run/grind_set_config.lean

Lines changed: 8 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -5,11 +5,18 @@ theorem fax : f (x + 1) = g (f x) := sorry
55

66
example : f (x + 100) = a := by
77
grind =>
8-
set_option gen 15 in
8+
set_config (gen := 15) -cutsat in
99
-- The following instantiations should not fail since we set
1010
-- `gen` to 15
1111
use [fax]; use [fax]; use [fax]; use [fax]; use [fax]
1212
use [fax]; use [fax]; use [fax]; use [fax]; use [fax]
1313
use [fax]; use [fax]; use [fax]; use [fax]; use [fax]
1414
fail_if_success use [fax] -- should fail
15+
fail_if_success have : 2*x ≠ 1 -- cutsat is disabled
16+
set_config +cutsat in
17+
have : 2*x ≠ 1
18+
set_config (cutsat := false) in
19+
fail_if_success have : 3*x ≠ 1
20+
set_config (cutsat := true) in
21+
have : 3*x ≠ 1
1522
sorry

0 commit comments

Comments
 (0)