Skip to content

Commit

Permalink
feat: No docs script
Browse files Browse the repository at this point in the history
  • Loading branch information
jstoobysmith committed Nov 11, 2024
1 parent 91d639c commit bff1468
Show file tree
Hide file tree
Showing 5 changed files with 114 additions and 0 deletions.
1 change: 1 addition & 0 deletions HepLean/AnomalyCancellation/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -57,6 +57,7 @@ instance chargesModule (χ : ACCSystemCharges) : Module ℚ χ.Charges :=
instance ChargesAddCommGroup (χ : ACCSystemCharges) : AddCommGroup χ.Charges :=
Module.addCommMonoidToAddCommGroup ℚ

/-- The module `χ.Charges` over `ℚ` is finite. -/
instance (χ : ACCSystemCharges) : Module.Finite ℚ χ.Charges :=
FiniteDimensional.finiteDimensional_pi ℚ

Expand Down
1 change: 1 addition & 0 deletions HepLean/AnomalyCancellation/GroupActions.lean
Original file line number Diff line number Diff line change
Expand Up @@ -37,6 +37,7 @@ structure ACCSystemGroupAction (χ : ACCSystem) where

namespace ACCSystemGroupAction

/-- The given instance of a group on the `group` field of a `ACCSystemGroupAction`. -/
instance {χ : ACCSystem} (G : ACCSystemGroupAction χ) : Group G.group := G.groupInst

/-- The action of a group element on the vector space of linear solutions. -/
Expand Down
3 changes: 3 additions & 0 deletions HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/ToSols.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,6 +39,7 @@ namespace AnomalyFreePerp
and for the cube line to sit in the quad. -/
def LineEqProp (R : MSSMACC.AnomalyFreePerp) : Prop := α₁ R = 0 ∧ α₂ R = 0 ∧ α₃ R = 0

/-- The proposition `LineEqProp` is decidable. -/
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (LineEqProp R) := instDecidableAnd

/-- A condition on `Sols` which we will show in `linEqPropSol_iff_proj_linEqProp` that is equivalent
Expand Down Expand Up @@ -79,6 +80,7 @@ entirely in the quadratic surface. -/
def InQuadProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
quadBiLin R.val R.val = 0 ∧ quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0

/-- The proposition `InQuadProp` is decidable. -/
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InQuadProp R) := instDecidableAnd

/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
Expand Down Expand Up @@ -124,6 +126,7 @@ def InCubeProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
cubeTriLin R.val R.val R.val = 0 ∧ cubeTriLin R.val R.val B₃.val = 0
cubeTriLin R.val R.val Y₃.val = 0

/-- The proposition `InCubeProp` is decidable. -/
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InCubeProp R) := instDecidableAnd

/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
Expand Down
4 changes: 4 additions & 0 deletions lakefile.toml
Original file line number Diff line number Diff line change
Expand Up @@ -64,6 +64,10 @@ srcDir = "scripts"
name = "free_simps"
srcDir = "scripts/MetaPrograms"

[[lean_exe]]
name = "no_docs"
srcDir = "scripts/MetaPrograms"

[[lean_exe]]
name = "informal"
supportInterpreter = true
Expand Down
105 changes: 105 additions & 0 deletions scripts/MetaPrograms/no_docs.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,105 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import Batteries.Lean.HashSet
import Lean
import Mathlib.Lean.Expr.Basic
import Mathlib.Lean.CoreM
import HepLean.Meta.Informal
import ImportGraph.RequiredModules
/-!
# Extracting commands with no doc strings.
-/

open Lean System Meta

def Imports.getConsts (imp : Import) : IO (Array ConstantInfo) := do
let mFile ← findOLean imp.module
let (modData, _) ← readModuleData mFile
pure modData.constants

def Imports.getUserConsts (imp : Import) : MetaM (Array ConstantInfo) := do
let env ← getEnv
let x := (← Imports.getConsts imp).filter (fun c => ¬ c.name.isInternal)
let x := x.filter (fun c => ¬ Lean.isCasesOnRecursor env c.name)
let x := x.filter (fun c => ¬ Lean.isRecOnRecursor env c.name)
let x := x.filter (fun c => ¬ Lean.isNoConfusion env c.name)
let x := x.filter (fun c => ¬ Lean.isBRecOnRecursor env c.name)
let x := x.filter (fun c => ¬ Lean.isAuxRecursorWithSuffix env c.name Lean.binductionOnSuffix)
let x := x.filter (fun c => ¬ Lean.isAuxRecursorWithSuffix env c.name Lean.belowSuffix)
let x := x.filter (fun c => ¬ Lean.isAuxRecursorWithSuffix env c.name Lean.ibelowSuffix)
/- Removing syntax category declarations. -/
let x := x.filter (fun c => ¬ c.name.toString = "Informal.informalAssignment.quot" )
let x := x.filter (fun c => ¬ c.name.toString = "TensorTree.indexExpr.quot" )
let x := x.filter (fun c => ¬ c.name.toString = "TensorTree.tensorExpr.quot" )
pure x

def Name.lineNumber (c : Name) : MetaM Nat := do
match ← findDeclarationRanges? c with
| some decl => pure decl.range.pos.line
| none => panic! s!"{c} is a declaration without position"

def Name.toFile (c : Name) : MetaM String := do
return s!"./{c.toString.replace "." "/" ++ ".lean"}"

def Name.location (c : Name) : MetaM String := do
let env ← getEnv
let x := env.getModuleFor? c
match x with
| some decl => pure ((← Name.toFile decl) ++ ":" ++ toString (← Name.lineNumber c) ++ " "
++ c.toString)
| none => panic! s!"{c} is a declaration without position"

def Name.hasPos (c : Name) : MetaM Bool := do
match ← findDeclarationRanges? c with
| some _ => pure true
| none => pure false

def Name.hasDocString (c : Name) : MetaM Bool := do
let env ← getEnv
let doc ← Lean.findDocString? env c
match doc with
| some _ => pure true
| none => pure false

def Imports.NoDocStringDef (imp : Import) : MetaM UInt32 := do
let x := (← Imports.getUserConsts imp).filter (fun c => c.isDef)
let x ← x.filterM (fun c => do
return Bool.not (← (Name.hasDocString c.name)))
let y ← x.filterM (fun c => Name.hasPos c.name)
let loc ← y.mapM (fun c => (Name.location c.name))
if loc.toList.length > 0 then
IO.println "\n"
IO.println s!"Module {imp.module} has the following definitions without doc strings:"
IO.println (String.intercalate "\n" loc.toList)
pure 0

def Imports.NoDocStringLemma (imp : Import) : MetaM UInt32 := do
let x := (← Imports.getUserConsts imp).filter (fun c => ¬ c.isDef)
let x ← x.filterM (fun c => do
return Bool.not (← (Name.hasDocString c.name)))
let y ← x.filterM (fun c => Name.hasPos c.name)
let loc ← y.mapM (fun c => (Name.location c.name))
if loc.toList.length > 0 then
IO.println "\n"
IO.println s!"Module {imp.module} has the following lemmas without doc strings:"
IO.println (String.intercalate "\n" loc.toList)
pure 0

unsafe def main (args : List String) : IO UInt32 := do
initSearchPath (← findSysroot)
let mods : Name := `HepLean
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 imports := hepLeanMod.imports.filter (fun c => c.module ≠ `Init)
let _ ← CoreM.withImportModules #[`HepLean] (imports.mapM Imports.NoDocStringDef).run'
if "--lemmas" ∈ args then
let _ ← CoreM.withImportModules #[`HepLean] (imports.mapM Imports.NoDocStringLemma).run'
pure 0

0 comments on commit bff1468

Please sign in to comment.