PhysLean/HepLean/Meta/Basic.lean

208 lines
7 KiB
Text
Raw Normal View History

2024-11-11 11:23:27 +00:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Lean.Expr.Basic
import ImportGraph.RequiredModules
/-!
## Basic Lean meta programming commands
-/
/-- The size of a flattened array of arrays. -/
def Array.flatSize {α} : Array (Array α) → Nat :=
foldl (init := 0) fun sizeAcc as => sizeAcc + as.size
/-- The size of a flattened array of arrays after applying an element-wise filter. -/
def Array.flatFilterSizeM {α m} [Monad m] (p : α → m Bool) : Array (Array α) → m Nat :=
foldlM (init := 0) fun sizeAcc as => return sizeAcc + (← as.filterM p).size
open Lean
2024-11-11 11:23:27 +00:00
/-!
## Imports
-/
/-- Gets all imports within HepLean. -/
def HepLean.allImports : IO (Array Import) := do
2024-11-11 11:23:27 +00:00
initSearchPath (← findSysroot)
let mods := `HepLean
let mFile ← findOLean mods
unless ← mFile.pathExists do
throw <| IO.userError s!"object file '{mFile}' of module {mods} does not exist"
2024-11-11 11:23:27 +00:00
let (hepLeanMod, _) ← readModuleData mFile
hepLeanMod.imports.filterM fun c => return c.module != `Init
2024-11-11 11:23:27 +00:00
/-- Number of files within HepLean. -/
def HepLean.noImports : IO Nat := do
2024-11-11 11:23:27 +00:00
let imports ← allImports
return imports.size
2024-11-11 11:23:27 +00:00
/-- Gets all constants from an import. -/
def HepLean.Imports.getConsts (imp : Import) : IO (Array ConstantInfo) := do
2024-11-11 11:23:27 +00:00
let mFile ← findOLean imp.module
let (modData, _) ← readModuleData mFile
return modData.constants
2024-11-11 11:23:27 +00:00
/-- Gets all user defined constants from an import. -/
def HepLean.Imports.getUserConsts (imp : Import) : CoreM (Array ConstantInfo) := do
2024-11-11 11:23:27 +00:00
let env ← getEnv
let consts ← Imports.getConsts imp
consts.filterM fun c =>
let name := c.name
return !name.isInternal
&& !isCasesOnRecursor env name
&& !isRecOnRecursor env name
&& !isNoConfusion env name
&& !isBRecOnRecursor env name
&& !isAuxRecursorWithSuffix env name binductionOnSuffix
&& !isAuxRecursorWithSuffix env name belowSuffix
&& !isAuxRecursorWithSuffix env name ibelowSuffix
/- Removing syntax category declarations. -/
&& name.toString != "TensorTree.indexExpr.quot"
&& name.toString != "TensorTree.tensorExpr.quot"
/-- Turns a name into a system file path. -/
def Lean.Name.toFilePath (c : Name) : System.FilePath :=
System.mkFilePath (c.toString.splitOn ".") |>.addExtension "lean"
2024-11-11 11:23:27 +00:00
/-- Lines from import. -/
def HepLean.Imports.getLines (imp : Import) : IO (Array String) := do
IO.FS.lines imp.module.toFilePath
2024-11-11 11:23:27 +00:00
namespace Lean.Name
2024-11-11 11:23:27 +00:00
/-!
## Name
-/
variable {m} [Monad m] [MonadEnv m] [MonadLiftT BaseIO m]
2024-11-11 11:23:27 +00:00
/-- Turns a name into a Lean file. -/
def toRelativeFilePath (c : Name) : System.FilePath :=
System.FilePath.join "." c.toFilePath
2024-11-11 11:23:27 +00:00
2024-12-05 07:12:56 +00:00
/-- Turns a name, which represents a module, into a link to github. -/
def toGitHubLink (c : Name) (line : Nat) : String :=
s!"https://github.com/HEPLean/HepLean/blob/master/{c.toFilePath}#L{line}"
2024-12-05 07:12:56 +00:00
2024-11-11 11:23:27 +00:00
/-- Given a name, returns the line number. -/
def lineNumber (c : Name) : m Nat := do
2024-11-12 10:00:23 +00:00
match ← findDeclarationRanges? c with
| some decl => return decl.range.pos.line
2024-11-11 11:23:27 +00:00
| none => panic! s!"{c} is a declaration without position"
2024-12-04 13:37:23 +00:00
/-- Given a name, returns the file name corresponding to that declaration. -/
def fileName (c : Name) : m Name := do
2024-12-04 13:37:23 +00:00
let env ← getEnv
match env.getModuleFor? c with
| some decl => return decl
2024-12-04 13:37:23 +00:00
| none => panic! s!"{c} is a declaration without position"
2024-11-11 11:23:27 +00:00
/-- Returns the location of a name. -/
def location (c : Name) : m String := do
2024-11-11 11:23:27 +00:00
let env ← getEnv
match env.getModuleFor? c with
| some decl => return s!"{decl.toRelativeFilePath}:{← c.lineNumber} {c}"
2024-11-11 11:23:27 +00:00
| none => panic! s!"{c} is a declaration without position"
/-- Determines if a name has a location. -/
def hasPos (c : Name) : m Bool := do
let ranges? ← findDeclarationRanges? c
return ranges?.isSome
2024-11-11 11:23:27 +00:00
/-- Determines if a name has a doc string. -/
def hasDocString (c : Name) : CoreM Bool := do
2024-11-11 11:23:27 +00:00
let env ← getEnv
let doc? ← findDocString? env c
return doc?.isSome
2024-11-11 11:23:27 +00:00
2025-01-23 14:31:03 +00:00
/-- The doc string from a name. -/
def getDocString (c : Name) : CoreM String := do
2025-01-23 10:46:50 +00:00
let env ← getEnv
let doc? ← findDocString? env c
return doc?.getD ""
2025-01-23 10:46:50 +00:00
/-- Given a name, returns the source code defining that name, including doc strings. -/
def getDeclString (name : Name) : CoreM String := do
2024-12-04 13:37:23 +00:00
let env ← getEnv
match ← findDeclarationRanges? name with
| some { range := { pos, endPos, .. }, .. } =>
match env.getModuleFor? name with
2024-12-04 16:24:23 +00:00
| some fileName =>
let fileContent ← IO.FS.readFile fileName.toRelativeFilePath
2024-12-04 13:37:23 +00:00
let fileMap := fileContent.toFileMap
return fileMap.source.extract (fileMap.ofPosition pos) (fileMap.ofPosition endPos)
| none => return ""
| none => return ""
/-- Given a name, returns the source code defining that name,
starting with the def ... or lemma... etc. -/
def getDeclStringNoDoc (name : Name) : CoreM String := do
let declerationString ← getDeclString name
let headerLine (line : String) : Bool :=
line.startsWith "def " line.startsWith "lemma " line.startsWith "inductive "
line.startsWith "structure " line.startsWith "theorem "
line.startsWith "instance " line.startsWith "abbrev "
line.startsWith "noncomputable def "
let lines := declerationString.splitOn "\n"
match lines.findIdx? headerLine with
| none => panic! s!"{name} has no header line"
| some i => return String.intercalate "\n" (lines.drop i)
end Lean.Name
namespace HepLean
2024-12-04 13:37:23 +00:00
2024-11-11 11:23:27 +00:00
/-- Number of definitions. -/
def noDefs : CoreM Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getUserConsts
x.flatFilterSizeM fun c => return c.isDef && (← c.name.hasPos)
2024-11-11 11:23:27 +00:00
/-- Number of definitions. -/
def noLemmas : CoreM Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getUserConsts
x.flatFilterSizeM fun c => return !c.isDef && (← c.name.hasPos)
2024-11-11 11:23:27 +00:00
/-- Number of definitions without a doc-string. -/
def noDefsNoDocString : CoreM Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getUserConsts
x.flatFilterSizeM fun c =>
return c.isDef && (← c.name.hasPos) && !(← c.name.hasDocString)
2024-11-11 11:23:27 +00:00
/-- Number of definitions without a doc-string. -/
def noLemmasNoDocString : CoreM Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getUserConsts
x.flatFilterSizeM fun c =>
return !c.isDef && (← c.name.hasPos) && !(← c.name.hasDocString)
2024-11-11 11:23:27 +00:00
/-- The number of lines in HepLean. -/
def noLines : IO Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getLines
return x.flatSize
2024-11-11 11:23:27 +00:00
2024-11-11 11:38:33 +00:00
/-- The number of TODO items. -/
def noTODOs : IO Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getLines
x.flatFilterSizeM fun l => return l.startsWith "TODO "
2024-11-11 11:38:33 +00:00
/-- The number of files with a TODO item. -/
def noFilesWithTODOs : IO Nat := do
let imports ← HepLean.allImports
let x ← imports.mapM HepLean.Imports.getLines
let x := x.filter fun bs => bs.any fun l => l.startsWith "TODO "
return x.size
2024-11-11 11:38:33 +00:00
2024-11-11 11:23:27 +00:00
end HepLean