Skip to content
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
5 changes: 5 additions & 0 deletions src/Init/Grind/Interactive.lean
Original file line number Diff line number Diff line change
Expand Up @@ -240,6 +240,11 @@ syntax (name := exposeNames) "expose_names" : grind
but it sets the option only within the tactics `tacs`. -/
syntax (name := setOption) "set_option " (ident (noWs "." noWs ident)?) ppSpace optionValue " in " grindSeq : grind

/--
`set_config configItem+ in tacs` executes `tacs` with the updated configuration options `configItem+`
-/
syntax (name := setConfig) "set_config " configItem+ " in " grindSeq : grind

/--
Proves `<term>` using the current `grind` state and default search strategy.
-/
Expand Down
32 changes: 9 additions & 23 deletions src/Lean/Elab/Tactic/Grind/BuiltinTactic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -426,30 +426,16 @@ where
liftGrindM <| resetAnchors
replaceMainGoal [{ goal with mvarId }]

def isGrindConfigField? (stx : Syntax) : CoreM (Option Name) := do
unless stx.isIdent do return none
let id := stx.getId
let env ← getEnv
let info := getStructureInfo env ``Grind.Config
unless info.fieldNames.contains id do return none
return some id

@[builtin_grind_tactic setOption] def elabSetOption : GrindTactic := fun stx => do
if let some fieldName ← isGrindConfigField? stx[1] then
let val := stx[3]
let val ← match val.isNatLit? with
| some num => pure <| DataValue.ofNat num
| none => match val with
| Syntax.atom _ "true" => pure <| DataValue.ofBool true
| Syntax.atom _ "false" => pure <| DataValue.ofBool false
| _ => throwErrorAt val "`grind` configuration option must be a Boolean or a numeral"
let config := (← read).ctx.config
let config ← setConfigField config fieldName val
withReader (fun c => { c with ctx.config := config }) do
evalGrindTactic stx[5]
else
let options ← Elab.elabSetOption stx[1] stx[3]
withOptions (fun _ => options) do evalGrindTactic stx[5]
let options ← Elab.elabSetOption stx[1] stx[3]
withOptions (fun _ => options) do evalGrindTactic stx[5]

@[builtin_grind_tactic setConfig] def elabSetConfig : GrindTactic := fun stx => do
let `(grind| set_config $[$items:configItem]* in $seq:grindSeq) := stx | throwUnsupportedSyntax
let config := (← read).ctx.config
let config ← elabConfigItems config items
withReader (fun c => { c with ctx.config := config }) do
evalGrindTactic seq

@[builtin_grind_tactic mbtc] def elabMBTC : GrindTactic := fun _ => do
liftGoalM do
Expand Down
16 changes: 16 additions & 0 deletions src/Lean/Elab/Tactic/Grind/Config.lean
Original file line number Diff line number Diff line change
Expand Up @@ -13,4 +13,20 @@ namespace Lean.Elab.Tactic.Grind
/-- Sets a field of the `grind` configuration object. -/
declare_config_getter setConfigField Grind.Config

def elabConfigItems (init : Grind.Config) (items : Array (TSyntax `Lean.Parser.Tactic.configItem))
: CoreM Grind.Config := do
let mut config := init
for item in items do
match item with
| `(Lean.Parser.Tactic.configItem| ($fieldName:ident := true))
| `(Lean.Parser.Tactic.configItem| +$fieldName:ident) =>
config ← setConfigField config fieldName.getId true
| `(Lean.Parser.Tactic.configItem| ($fieldName:ident := false))
| `(Lean.Parser.Tactic.configItem| -$fieldName:ident) =>
config ← setConfigField config fieldName.getId false
| `(Lean.Parser.Tactic.configItem| ($fieldName:ident := $val:num)) =>
config ← setConfigField config fieldName.getId (.ofNat val.getNat)
| _ => throwErrorAt item "unexpected configuration option"
return config

end Lean.Elab.Tactic.Grind
9 changes: 8 additions & 1 deletion tests/lean/run/grind_set_config.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,11 +5,18 @@ theorem fax : f (x + 1) = g (f x) := sorry

example : f (x + 100) = a := by
grind =>
set_option gen 15 in
set_config (gen := 15) -cutsat in
-- The following instantiations should not fail since we set
-- `gen` to 15
use [fax]; use [fax]; use [fax]; use [fax]; use [fax]
use [fax]; use [fax]; use [fax]; use [fax]; use [fax]
use [fax]; use [fax]; use [fax]; use [fax]; use [fax]
fail_if_success use [fax] -- should fail
fail_if_success have : 2*x ≠ 1 -- cutsat is disabled
set_config +cutsat in
have : 2*x ≠ 1
set_config (cutsat := false) in
fail_if_success have : 3*x ≠ 1
set_config (cutsat := true) in
have : 3*x ≠ 1
sorry
Loading