PhysLean/scripts/find_TODOs.lean

99 lines
3 KiB
Text
Raw Normal View History

2024-07-09 15:25:06 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import Lean
import Batteries.Data.String.Matcher
2024-11-04 06:56:21 +00:00
import Mathlib.Data.Nat.Notation
2024-07-09 15:25:06 -04:00
/-!
# TODO finder
2025-02-14 08:45:02 +00:00
This program finds all instances of `/<!> TODO: ...` (without the `<>`) in PhysLean files.
2024-07-09 15:25:06 -04:00
## Note
Parts of this file are adapted from `Mathlib.Tactic.Linter.TextBased`,
authored by Michael Rothgang.
-/
open Lean System Meta
/-- Given a list of lines, outputs an error message and a line number. -/
2025-02-14 08:45:02 +00:00
def PhysLeanTODOItem : Type := Array String → Array (String × )
2024-07-09 15:25:06 -04:00
/-- Checks if a . -/
2025-02-14 08:45:02 +00:00
def TODOFinder : PhysLeanTODOItem := fun lines ↦ Id.run do
2024-07-09 15:25:06 -04:00
let enumLines := (lines.toList.enumFrom 1)
let todos := enumLines.filterMap (fun (lno1, l1) ↦
if l1.startsWith "/-! TODO:" then
some ((l1.replace "/-! TODO: " "").replace "-/" "", lno1)
else none)
todos.toArray
structure TODOContext where
/-- The underlying `message`. -/
statement : String
/-- The line number -/
lineNumber :
/-- The file path -/
path : FilePath
def printTODO (todos : Array TODOContext) : IO Unit := do
for e in todos do
IO.println (s!"{e.path}:{e.lineNumber}: {e.statement}")
def filePathToGitPath (S : FilePath) (n : ) : String :=
2025-02-14 08:45:02 +00:00
"https://github.com/HEPLean/PhysLean/blob/master/"++
2024-07-09 15:25:06 -04:00
(S.toString.replace "." "/").replace "/lean" ".lean"
++ "#L" ++ toString n
def docTODO (todos : Array TODOContext) : IO String := do
let mut out := ""
for e in todos do
out := out ++ (s!" - [{e.statement}]("++ filePathToGitPath e.path e.lineNumber ++ ")\n")
return out
def hepLeanLintFile (path : FilePath) : IO String := do
let lines ← IO.FS.lines path
let allOutput := (Array.map (fun lint ↦
(Array.map (fun (e, n) ↦ TODOContext.mk e n path)) (lint lines)))
#[TODOFinder]
let errors := allOutput.flatten
printTODO errors
docTODO errors
def todoFileHeader : String := s!"
# TODO List
2025-02-14 08:45:02 +00:00
This is an automatically generated list of TODOs appearing as `/-! TODO:...` in PhysLean.
2024-07-09 15:25:06 -04:00
Please feel free to contribute to the completion of these tasks.
"
def main (args : List String) : IO UInt32 := do
initSearchPath (← findSysroot)
2025-02-14 08:45:02 +00:00
let mods : Name := `PhysLean
2024-07-09 15:25:06 -04:00
let imp : Import := {module := mods}
let mFile ← findOLean imp.module
unless (← mFile.pathExists) do
throw <| IO.userError s!"object file '{mFile}' of module {imp.module} does not exist"
let (hepLeanMod, _) ← readModuleData mFile
let mut out : String := ""
for imp in hepLeanMod.imports do
if imp.module == `Init then continue
let filePath := (mkFilePath (imp.module.toString.split (· == '.'))).addExtension "lean"
let l ← hepLeanLintFile filePath
if l != "" then
out := out ++ "\n### " ++ imp.module.toString ++ "\n"
out := out ++ l
let fileOut : System.FilePath := {toString := "./docs/TODOList.md"}
/- Below here is concerned with writing out to the docs. -/
if "mkFile" ∈ args then
IO.println (s!"TODOList file made.")
IO.FS.writeFile fileOut (todoFileHeader ++ out)
return 0