-
Notifications
You must be signed in to change notification settings - Fork 19
/
ProblemExtraction.lean
375 lines (303 loc) · 13.2 KB
/
ProblemExtraction.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
import Lean.Elab.Command
import Lean.Elab.Eval
import Lean.Meta.Basic
import Batteries.Data.String.Basic
import Batteries.Lean.NameMapAttribute
import Lean
/-!
Special commands to aid in "problem extraction".
For the math problems that we archive, we aim to include proofs in-line.
Sometimes, however, we want to present the problems without giving away
information about the solutions.
Therefore, we have "problem extraction" -- a means of stripping solutions.
During problem extraction, all declarations are removed
except those that have been tagged with one of the below command wrappers.
-/
namespace ProblemExtraction
open Lean Elab
structure Replacement where
startPos : String.Pos
endPos : String.Pos
newValue : String
deriving Inhabited
inductive EntryVariant where
/-- full file text and the position where extraction should start-/
| file : String → String.Pos → EntryVariant
/-- substring replacement. positions are relative to the full file -/
| replace : Replacement → EntryVariant
| snip_begin : String.Pos → EntryVariant
| snip_end : String.Pos → EntryVariant
/-- An entry in the state of the Problem Extraction environment extension -/
structure Entry where
/-- The module where the entry originated. -/
(module : Name)
/-- Lean code to be included in the extracted problem file. -/
(variant : EntryVariant)
abbrev ExtractionExtension := SimplePersistentEnvExtension Entry (Array Entry)
initialize problemExtractionExtension : ExtractionExtension ←
registerSimplePersistentEnvExtension {
name := `problem_extraction
addImportedFn := Array.concatMap id
addEntryFn := Array.push
}
initialize solutionExtractionExtension : ExtractionExtension ←
registerSimplePersistentEnvExtension {
name := `solution_extraction
addImportedFn := Array.concatMap id
addEntryFn := Array.push
}
abbrev DetermineDeclsExtension := SimplePersistentEnvExtension Name (Array Name)
initialize determineDeclsExtension : DetermineDeclsExtension ←
registerSimplePersistentEnvExtension {
name := `determine_decls
addImportedFn := Array.concatMap id
addEntryFn := Array.push
}
inductive ProblemTag where
| Algebra : ProblemTag
| NumberTheory : ProblemTag
| Combinatorics : ProblemTag
| Geometry : ProblemTag
| Inequality : ProblemTag
deriving Ord
def ProblemTag.toNat (t : ProblemTag) : Nat := match t with
| .Algebra => 0
| .NumberTheory => 1
| .Combinatorics => 2
| .Geometry => 3
| .Inequality => 4
instance : ToString ProblemTag where
toString := fun p => match p with
| .Geometry => "geometry"
| .Inequality => "inequality"
| .Combinatorics => "combinatorics"
| .NumberTheory => "number theory"
| .Algebra => "algebra"
structure ProblemFileMetadata where
tags : List ProblemTag := []
--- If the formalized solution was imported from somewhere else,
--- then this field should contain the URL of that source.
importedFrom : Option String := .none
--- Names of the people who wrote the solution. By default, this
--- is automatically populated via the file's copyright header.
authors : List String := []
--- Names of the people who wrote the solution. This field
--- is automatically populated via the file's copyright header,
--- which is assumed to be everything before the first 'import'.
copyrightHeader : String := ""
structure ProblemMetadataEntry where
module : Name
metadata : ProblemFileMetadata
abbrev ProblemMetadataExtension :=
SimplePersistentEnvExtension ProblemMetadataEntry (Array ProblemMetadataEntry)
initialize problemMetadataExtension : ProblemMetadataExtension ←
registerSimplePersistentEnvExtension {
name := `problem_metadata
addImportedFn := Array.concatMap id
addEntryFn := Array.push
}
def parseAuthors (src : String) : List String := Id.run do
let lines := src.splitOn "\n"
for l in lines do
if l.startsWith "Authors: "
then
return (l.stripPrefix "Authors: ").split (fun c => c = ',')
return []
def parseCopyrightHeader (src : String) : String := Id.run do
let lines := src.splitOn "\n"
let mut result := ""
for l in lines do
if l.startsWith "import"
then break
else
result := result ++ l ++ "\n"
return result
/-- Top-level command to mark that a file should participate in problem extraction.
This should be at the top of the file (after imports); content above it is ignored
during problem extraction (except for imports). -/
syntax (name := problemFile) "problem_file " (term)? : command
def elabProblemFile (tk : Syntax) (md : Option (TSyntax `term)) : Command.CommandElabM Unit := do
let .some startPos := (match md with
| .some md => md.raw.getTailPos?
| .none => tk.getTailPos?) | throwError "problem_file syntax has no tail pos"
let src := (←read).fileMap.source
let startPos := ⟨startPos.byteIdx + 1⟩ -- HACK: add one to consume unwanted newline
let mod := (←getEnv).header.mainModule
modifyEnv fun env =>
problemExtractionExtension.addEntry env ⟨mod, EntryVariant.file src startPos⟩
modifyEnv fun env =>
solutionExtractionExtension.addEntry env ⟨mod, EntryVariant.file src startPos⟩
let mut mdv ← match md with
| some stx => Lean.Elab.Command.liftTermElabM do
unsafe Lean.Elab.Term.evalTerm ProblemFileMetadata (mkConst ``ProblemFileMetadata) stx
| .none => pure {}
let mdv' :=
if mdv.authors.isEmpty
then { mdv with authors := parseAuthors src }
else mdv
let mdv' := { mdv' with copyrightHeader := parseCopyrightHeader src }
modifyEnv fun env => problemMetadataExtension.addEntry env ⟨mod, mdv'⟩
elab_rules : command
| `(command| problem_file%$tk) => elabProblemFile tk none
| `(command| problem_file%$tk $md) => elabProblemFile tk md
/-- Starts a group of commands that will be discarded by problem extraction. -/
syntax (name := snipBegin) "snip begin" : command
/-- Ends a group of commands that will be discarded by problem extraction. -/
syntax (name := snipEnd) "snip end" : command
elab_rules : command
| `(command| snip begin%$tk1) => do
let .some startPos := tk1.getPos? | throwError "snip syntax has no start pos"
let .some endPos := tk1.getTailPos? | throwError "snip syntax has no tail pos"
let startPos := ⟨startPos.byteIdx - 1⟩ -- HACK: subtract one to consume unwanted newline
let mod := (←getEnv).header.mainModule
let ext := problemExtractionExtension
modifyEnv fun env => ext.addEntry env ⟨mod, EntryVariant.snip_begin startPos⟩
modifyEnv fun env => solutionExtractionExtension.addEntry env
⟨mod, EntryVariant.replace ⟨startPos, endPos, ""⟩⟩
| `(command| snip end%$tk2) => do
let .some startPos := tk2.getPos? | throwError "snip syntax has no start pos"
let .some endPos := tk2.getTailPos? | throwError "snip syntax has no end pos"
let endPos := ⟨endPos.byteIdx + 1⟩ -- HACK: add one to consume unwanted newline
let mod := (←getEnv).header.mainModule
let ext := problemExtractionExtension
modifyEnv fun env => ext.addEntry env ⟨mod, EntryVariant.snip_end endPos⟩
modifyEnv fun env => solutionExtractionExtension.addEntry env
⟨mod, EntryVariant.replace ⟨startPos, endPos, ""⟩⟩
/--
A synonym for `theorem`. Indicates that a declaration is a problem statement.
During problem extraction, the proof is replaced by a `sorry`.
-/
syntax (name := problem) declModifiers "problem " declId ppIndent(declSig) declVal : command
elab_rules : command
| `(command| $dm:declModifiers problem%$pb $di:declId $ds:declSig $dv:declVal) => do
let mod := (←getEnv).header.mainModule
let (.some pStartPos, .some pEndPos) := (pb.getPos?, pb.getTailPos?)
| throwError "failed to get problem syntax"
modifyEnv fun env => problemExtractionExtension.addEntry env ⟨mod,
EntryVariant.replace ⟨pStartPos, pEndPos, "theorem"⟩⟩
modifyEnv fun env => solutionExtractionExtension.addEntry env ⟨mod,
EntryVariant.replace ⟨pStartPos, pEndPos, "theorem"⟩⟩
let (.some vStartPos, .some vEndPos) := (dv.raw.getPos?, dv.raw.getTailPos?)
| throwError "failed to get declVal syntax"
modifyEnv fun env => problemExtractionExtension.addEntry env ⟨mod,
EntryVariant.replace ⟨vStartPos, vEndPos, ":= sorry"⟩⟩
let cmd ← `(command | $dm:declModifiers theorem $di $ds $dv)
Lean.Elab.Command.elabCommand cmd
/--
A synonym for `abbrev`. Marks data that is intended to be filled in as part of
a solution. During problem extraction, the body of the decl is replaced by a `sorry`.
During judging, a human will inspect the filled-in body
to see whether it is reasonable.
-/
syntax (name := determine)
declModifiers "determine " declId ppIndent(optDeclSig) declVal : command
elab_rules : command
| `(command| $dm:declModifiers determine%$dt $di:declId $ds:optDeclSig $dv:declVal) => do
let mod := (←getEnv).header.mainModule
let (.some dStartPos, .some dEndPos) := (dt.getPos?, dt.getTailPos?)
| throwError "failed to get problem syntax"
modifyEnv fun env => problemExtractionExtension.addEntry env ⟨mod,
EntryVariant.replace ⟨dStartPos, dEndPos, "/- determine -/ abbrev"⟩⟩
modifyEnv fun env => solutionExtractionExtension.addEntry env ⟨mod,
EntryVariant.replace ⟨dStartPos, dEndPos, "/- determine -/ abbrev"⟩⟩
let (.some vStartPos, .some vEndPos) := (dv.raw.getPos?, dv.raw.getTailPos?)
| throwError "failed to get declVal syntax"
modifyEnv fun env => problemExtractionExtension.addEntry env ⟨mod,
EntryVariant.replace ⟨vStartPos, vEndPos, ":= sorry"⟩⟩
let cmd ← `(command | set_option linter.unusedVariables false in $dm:declModifiers abbrev $di:declId $ds:optDeclSig $dv:declVal)
Lean.Elab.Command.elabCommand cmd
match di with
| `(Lean.Parser.Command.declId | $i:ident) =>
let name ← Lean.resolveGlobalConstNoOverload i
modifyEnv fun env => determineDeclsExtension.addEntry env name
| _ => throwError "explicit universes in `determine` are currently unsupported"
/--
Prints the current contents of the Problem Extraction extension.
-/
syntax (name := showExtraction) "#show_problem_extraction" : command
elab_rules : command
| `(command| #show_problem_extraction) => do
let ext := problemExtractionExtension
let env ← getEnv
let st := ext.getState env
IO.println s!"ProblemExtraction st.size = {st.size}"
for ⟨filename, _⟩ in st do
IO.println s!"{filename}"
let st := determineDeclsExtension.getState env
IO.println s!"Determine decls:"
for n in st do
IO.println s!"{n}"
/--
Helper function for extractProblems.
-/
private def findModuleImports
{m : Type → Type} [Monad m] [MonadError m] (env : Environment) (md : Name) :
m (Array Import) := do
let moduleNames := env.header.moduleNames
let mut idx := 0
for m1 in moduleNames do
if m1 = md
then return (env.header.moduleData.get! idx).imports
else idx := 1 + idx
throwError s!"module {md} not found"
def extractFromExt {m : Type → Type} [Monad m] [MonadEnv m] [MonadError m]
(ext : ExtractionExtension) : m (NameMap String) := do
let env ← getEnv
let st := ext.getState env
let mut inProgress : NameMap (String × String.Pos × String) := mkNameMap _
for ⟨module, variant⟩ in st do
match variant with
| .file s p =>
inProgress := inProgress.insert module ⟨s, p, ""⟩
| .replace ⟨startPos, endPos, s⟩ =>
match inProgress.find? module with
| .some ⟨src, cur, acc⟩ =>
inProgress := inProgress.insert module
⟨src, endPos, acc ++ (Substring.mk src cur startPos).toString ++ s⟩
| .none => pure ()
| .snip_begin pos =>
match inProgress.find? module with
| .some ⟨src, cur, acc⟩ =>
inProgress := inProgress.insert module
⟨src, pos, acc ++ (Substring.mk src cur pos).toString⟩
| .none => pure ()
| .snip_end pos =>
match inProgress.find? module with
| .some ⟨src, _, acc⟩ =>
inProgress := inProgress.insert module ⟨src, pos, acc⟩
| .none => pure ()
let mut result := mkNameMap _
for ⟨module, ⟨src, endPos, acc⟩⟩ in inProgress do
let mut imports := ""
for im in ← findModuleImports env module do
if im.module.toString ≠ "Init" && im.module ≠ `ProblemExtraction
then imports := imports ++ s!"import {im}\n"
result := result.insert module
(imports ++ acc ++ (Substring.mk src endPos src.endPos).toString)
pure result
/--
Using the data in the problem extraction environment extension,
constructs a map from module name to problem source code.
-/
def extractProblems {m : Type → Type} [Monad m] [MonadEnv m] [MonadError m] :
m (NameMap String) :=
extractFromExt problemExtractionExtension
/--
Using the data in the solution extraction environment extension,
constructs a map from module name to solution source code.
-/
def extractSolutions {m : Type → Type} [Monad m] [MonadEnv m] [MonadError m] :
m (NameMap String) :=
extractFromExt solutionExtractionExtension
/--
Using the data in the solution extraction environment extension,
constructs a map from module name to problem metadata
-/
def extractMetadata {m : Type → Type} [Monad m] [MonadEnv m] [MonadError m] :
m (NameMap ProblemFileMetadata) := do
let env ← getEnv
let st := problemMetadataExtension.getState env
let mut result := mkNameMap _
for ⟨module, md⟩ in st do
result := result.insert module md
return result