@@ -132,6 +132,47 @@ def process (input : String) (env : Environment) (opts : Options) (fileName : Op
132
132
let s ← IO.processCommands inputCtx { : Parser.ModuleParserState } (Command.mkState env {} opts)
133
133
pure (s.commandState.env, s.commandState.messages)
134
134
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
+
135
176
@[export lean_run_frontend]
136
177
def runFrontend
137
178
(input : String)
@@ -151,6 +192,8 @@ def runFrontend
151
192
let (header, parserState, messages) ← Parser.parseHeader inputCtx
152
193
-- allow `env` to be leaked, which would live until the end of the process anyway
153
194
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
154
197
let env := env.setMainModule mainModuleName
155
198
let mut commandState := Command.mkState env messages opts
156
199
let elabStartTime := (← IO.monoNanosNow).toFloat / 1000000000
0 commit comments