Skip to content

Commit b07384a

Browse files
authored
feat: accept user-defined options on the cmdline (#4741)
Initial options are now re-parsed and validated after importing. Cmdline option assignments prefixed with `weak.` are silently discarded if the option name without the prefix does not exist. Fixes #3403
1 parent efc99b9 commit b07384a

File tree

5 files changed

+67
-6
lines changed

5 files changed

+67
-6
lines changed

src/Lean/Elab/Frontend.lean

Lines changed: 43 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -132,6 +132,47 @@ def process (input : String) (env : Environment) (opts : Options) (fileName : Op
132132
let s ← IO.processCommands inputCtx { : Parser.ModuleParserState } (Command.mkState env {} opts)
133133
pure (s.commandState.env, s.commandState.messages)
134134

135+
/--
136+
Parses values of options registered during import and left by the C++ frontend as strings, fails if
137+
any option names remain unknown.
138+
-/
139+
def reparseOptions (opts : Options) : IO Options := do
140+
let mut opts := opts
141+
let decls ← getOptionDecls
142+
for (name, val) in opts do
143+
let .ofString val := val
144+
| continue -- Already parsed by C++
145+
-- Options can be prefixed with `weak` in order to turn off the error when the option is not
146+
-- defined
147+
let weak := name.getRoot == `weak
148+
if weak then
149+
opts := opts.erase name
150+
let name := name.replacePrefix `weak Name.anonymous
151+
let some decl := decls.find? name
152+
| unless weak do
153+
throw <| .userError s!"invalid -D parameter, unknown configuration option '{name}'
154+
155+
If the option is defined in this library, use '-D{`weak ++ name}' to set it conditionally"
156+
157+
match decl.defValue with
158+
| .ofBool _ =>
159+
match val with
160+
| "true" => opts := opts.insert name true
161+
| "false" => opts := opts.insert name false
162+
| _ =>
163+
throw <| .userError s!"invalid -D parameter, invalid configuration option '{val}' value, \
164+
it must be true/false"
165+
| .ofNat _ =>
166+
let some val := val.toNat?
167+
| throw <| .userError s!"invalid -D parameter, invalid configuration option '{val}' value, \
168+
it must be a natural number"
169+
opts := opts.insert name val
170+
| .ofString _ => opts := opts.insert name val
171+
| _ => throw <| .userError s!"invalid -D parameter, configuration option '{name}' \
172+
cannot be set in the command line, use set_option command"
173+
174+
return opts
175+
135176
@[export lean_run_frontend]
136177
def runFrontend
137178
(input : String)
@@ -151,6 +192,8 @@ def runFrontend
151192
let (header, parserState, messages) ← Parser.parseHeader inputCtx
152193
-- allow `env` to be leaked, which would live until the end of the process anyway
153194
let (env, messages) ← processHeader (leakEnv := true) header opts messages inputCtx trustLevel
195+
-- now that imports have been loaded, check options again
196+
let opts ← reparseOptions opts
154197
let env := env.setMainModule mainModuleName
155198
let mut commandState := Command.mkState env messages opts
156199
let elabStartTime := (← IO.monoNanosNow).toFloat / 1000000000

src/util/shell.cpp

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -310,7 +310,9 @@ options set_config_option(options const & opts, char const * in) {
310310
<< "' cannot be set in the command line, use set_option command");
311311
}
312312
} else {
313-
throw lean::exception(lean::sstream() << "invalid -D parameter, unknown configuration option '" << opt << "'");
313+
// More options may be registered by imports, so we leave validating them to the Lean side.
314+
// This (minor) duplication will be resolved when this file is rewritten in Lean.
315+
return opts.update(opt, val.c_str());
314316
}
315317
}
316318

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,12 @@
1+
import UserOpt.Opts
2+
3+
/-! Test setting user options from lakefile. -/
4+
5+
open Lean
6+
7+
def tst2 : MetaM Unit := do
8+
assert! myBoolOpt.get (← getOptions)
9+
assert! myNatOpt.get (← getOptions) == 4
10+
pure ()
11+
12+
#eval tst2

tests/pkg/user_opt/lakefile.lean

Lines changed: 0 additions & 5 deletions
This file was deleted.

tests/pkg/user_opt/lakefile.toml

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,9 @@
1+
name = "user_opt"
2+
defaultTargets = ["UserOpt", "UserOptCmdline"]
3+
4+
[[lean_lib]]
5+
name = "UserOpt"
6+
7+
[[lean_lib]]
8+
name = "UserOptCmdline"
9+
leanOptions = { myBoolOpt = true, weak.myNatOpt = 4, weak.notMyNatOpt = 5 }

0 commit comments

Comments
 (0)