/- Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved. Released under Apache 2.0 license. Authors: Joseph Tooby-Smith -/ import HepLean.Meta.Basic /-! # HepLean Stats This file concerns with statistics of HepLean. -/ open Lean System Meta HepLean def getStats : MetaM String := do let noDefsVal ← noDefs let noLemmasVal ← noLemmas let noImportsVal ← noImports let noDefsNoDocVal ← noDefsNoDocString let noLemmasNoDocVal ← noLemmasNoDocString let noLinesVal ← noLines let s := s!" Number of Imports: {noImportsVal} Number of Definitions: {noDefsVal} Number of Lemmas: {noLemmasVal} Number of Definitions without doc strings: {noDefsNoDocVal} (out of {noDefsVal}) Number of Lemmas without doc strings: {noLemmasNoDocVal} (out of {noLemmasVal}) Number of Lines: {noLinesVal}" pure s unsafe def Stats.toHtml : MetaM String := do let noDefsVal ← noDefs let noLemmasVal ← noLemmas let noImportsVal ← noImports let noDefsNoDocVal ← noDefsNoDocString let noLemmasNoDocVal ← noLemmasNoDocString let noLinesVal ← noLines let noInformalDefsVal ← noInformalDefs let noInformalLemmasVal ← noInformalLemmas let header := "--- layout: default ---
- Of which {noDefsVal - noDefsNoDocVal- noInformalLemmasVal} have doc-strings:
- Of which {noDefsVal - noInformalLemmasVal - noInformalDefsVal} are not informal definitions:
- Of which {noLemmasVal - noLemmasNoDocVal + noInformalLemmasVal} have doc-strings:
- Of which {noLemmasVal} are not informal lemmas:
" let footer := " " pure (header ++ "\n" ++ body ++ "\n" ++ footer) unsafe def main (args : List String) : IO UInt32 := do let _ ← noImports let statString ← CoreM.withImportModules #[`HepLean] (getStats).run' println! statString if "mkHTML" ∈ args then let html ← CoreM.withImportModules #[`HepLean] (Stats.toHtml).run' let htmlFile : System.FilePath := {toString := "./docs/Stats.html"} IO.FS.writeFile htmlFile html IO.println (s!"HTML file made.") pure 0