2024-12-05 06:49:50 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.Meta.Basic
|
2025-02-02 03:17:17 +08:00
|
|
|
|
import HepLean.Meta.Informal.Basic
|
2024-12-05 06:49:50 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## Informal definitions and lemmas
|
|
|
|
|
|
|
|
|
|
-/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
open Lean
|
2024-12-05 06:49:50 +00:00
|
|
|
|
|
|
|
|
|
namespace Informal
|
|
|
|
|
|
|
|
|
|
/-- Is true if and only if a `ConstantInfo` corresponds to an `InformalLemma` or a
|
|
|
|
|
`InformalDefinition`. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
def isInformal : ConstantInfo → Bool
|
|
|
|
|
| .defnInfo c => c.type.isConstOf ``InformalDefinition ∨ c.type.isConstOf ``InformalLemma
|
2024-12-05 06:49:50 +00:00
|
|
|
|
| _ => false
|
|
|
|
|
|
|
|
|
|
/-- Is true if and only if a `ConstantInfo` corresponds to an `InformalLemma`. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
def isInformalLemma : ConstantInfo → Bool
|
|
|
|
|
| .defnInfo c => c.type.isConstOf ``InformalLemma
|
2024-12-05 06:49:50 +00:00
|
|
|
|
| _ => false
|
|
|
|
|
|
|
|
|
|
/-- Is true if and only if a `ConstantInfo` corresponds to an `InformalDefinition`. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
def isInformalDef : ConstantInfo → Bool
|
|
|
|
|
| .defnInfo c => c.type.isConstOf ``InformalDefinition
|
2024-12-05 06:49:50 +00:00
|
|
|
|
| _ => false
|
|
|
|
|
|
|
|
|
|
/-- Takes a `ConstantInfo` corresponding to a `InformalLemma` and returns
|
|
|
|
|
the corresponding `InformalLemma`. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
unsafe def constantInfoToInformalLemma (c : ConstantInfo) : CoreM InformalLemma := do
|
2024-12-05 06:49:50 +00:00
|
|
|
|
match c with
|
2025-02-02 03:17:17 +08:00
|
|
|
|
| .defnInfo c => evalConstCheck InformalLemma ``InformalLemma c.name
|
2024-12-05 06:49:50 +00:00
|
|
|
|
| _ => panic! "Passed constantInfoToInformalLemma a `ConstantInfo` that is not a `InformalLemma`"
|
|
|
|
|
|
|
|
|
|
/-- Takes a `ConstantInfo` corresponding to a `InformalDefinition` and returns
|
|
|
|
|
the corresponding `InformalDefinition`. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
unsafe def constantInfoToInformalDefinition (c : ConstantInfo) : CoreM InformalDefinition := do
|
2024-12-05 06:49:50 +00:00
|
|
|
|
match c with
|
2025-02-02 03:17:17 +08:00
|
|
|
|
| .defnInfo c => evalConstCheck InformalDefinition ``InformalDefinition c.name
|
|
|
|
|
| _ => panic!
|
|
|
|
|
"Passed constantInfoToInformalDefinition a `ConstantInfo` that is not a `InformalDefinition`"
|
2024-12-05 06:49:50 +00:00
|
|
|
|
|
|
|
|
|
end Informal
|
|
|
|
|
|
|
|
|
|
namespace HepLean
|
|
|
|
|
|
|
|
|
|
/-- The number of informal lemmas in HepLean. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
def noInformalLemmas : CoreM Nat := do
|
2024-12-05 06:49:50 +00:00
|
|
|
|
let imports ← allImports
|
|
|
|
|
let x ← imports.mapM Imports.getUserConsts
|
2025-02-02 03:17:17 +08:00
|
|
|
|
x.flatFilterSizeM fun c => return Informal.isInformalLemma c
|
2024-12-05 06:49:50 +00:00
|
|
|
|
|
|
|
|
|
/-- The number of informal definitions in HepLean. -/
|
2025-02-02 03:17:17 +08:00
|
|
|
|
def noInformalDefs : CoreM Nat := do
|
2024-12-05 06:49:50 +00:00
|
|
|
|
let imports ← allImports
|
|
|
|
|
let x ← imports.mapM Imports.getUserConsts
|
2025-02-02 03:17:17 +08:00
|
|
|
|
x.flatFilterSizeM fun c => return Informal.isInformalDef c
|
2024-12-05 06:49:50 +00:00
|
|
|
|
|
|
|
|
|
end HepLean
|