2023-12-05 19:07:00 -08:00
|
|
|
|
import LSpec
|
|
|
|
|
import Pantograph.Serial
|
2023-12-15 10:40:36 -08:00
|
|
|
|
import Pantograph.Environment
|
2024-01-16 13:29:30 -08:00
|
|
|
|
import Test.Common
|
2024-03-28 18:56:42 -07:00
|
|
|
|
import Lean
|
2023-12-05 19:07:00 -08:00
|
|
|
|
|
2024-03-28 18:56:42 -07:00
|
|
|
|
open Lean
|
2023-12-26 09:22:57 -08:00
|
|
|
|
namespace Pantograph.Test.Environment
|
2023-12-05 19:07:00 -08:00
|
|
|
|
|
|
|
|
|
open Pantograph
|
|
|
|
|
|
2024-01-16 13:29:30 -08:00
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.InductInfo
|
2024-01-16 14:11:52 -08:00
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.ConstructorInfo
|
2024-04-09 21:24:08 -07:00
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.RecursorRule
|
2024-01-16 14:11:52 -08:00
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.RecursorInfo
|
2024-01-16 13:29:30 -08:00
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.EnvInspectResult
|
|
|
|
|
|
2024-03-28 19:56:03 -07:00
|
|
|
|
def test_catalog: IO LSpec.TestSeq := do
|
|
|
|
|
let env: Environment ← importModules
|
|
|
|
|
(imports := #[`Init])
|
|
|
|
|
(opts := {})
|
|
|
|
|
(trustLevel := 1)
|
|
|
|
|
let inner: CoreM LSpec.TestSeq := do
|
|
|
|
|
let catalogResult ← Environment.catalog {}
|
|
|
|
|
let symbolsWithNum := env.constants.fold (init := #[]) (λ acc name info =>
|
|
|
|
|
match (Environment.toFilteredSymbol name info).isSome && (name matches .num _ _) with
|
|
|
|
|
| false => acc
|
|
|
|
|
| true => acc.push name
|
|
|
|
|
)
|
|
|
|
|
return LSpec.check "No num symbols" (symbolsWithNum.size == 0)
|
|
|
|
|
runCoreMSeq env inner
|
|
|
|
|
|
2024-03-28 19:49:44 -07:00
|
|
|
|
def test_symbol_visibility: IO LSpec.TestSeq := do
|
2023-12-05 19:07:00 -08:00
|
|
|
|
let entries: List (Name × Bool) := [
|
|
|
|
|
("Nat.add_comm".toName, false),
|
2024-10-04 12:58:16 -07:00
|
|
|
|
("foo.bla.Init.Data.List.Basic.2.1.Init.Lean.Expr._hyg.4".toName, true),
|
2024-03-28 19:49:44 -07:00
|
|
|
|
("Init.Data.Nat.Basic._auxLemma.4".toName, true),
|
2023-12-05 19:07:00 -08:00
|
|
|
|
]
|
|
|
|
|
let suite := entries.foldl (λ suites (symbol, target) =>
|
2024-03-28 19:49:44 -07:00
|
|
|
|
let test := LSpec.check symbol.toString ((Environment.isNameInternal symbol) == target)
|
2023-12-05 19:07:00 -08:00
|
|
|
|
LSpec.TestSeq.append suites test) LSpec.TestSeq.done
|
|
|
|
|
return suite
|
|
|
|
|
|
2024-01-16 14:11:52 -08:00
|
|
|
|
inductive ConstantCat where
|
|
|
|
|
| induct (info: Protocol.InductInfo)
|
|
|
|
|
| ctor (info: Protocol.ConstructorInfo)
|
|
|
|
|
| recursor (info: Protocol.RecursorInfo)
|
|
|
|
|
|
2024-03-28 19:56:03 -07:00
|
|
|
|
def test_inspect: IO LSpec.TestSeq := do
|
|
|
|
|
let env: Environment ← importModules
|
|
|
|
|
(imports := #[`Init])
|
|
|
|
|
(opts := {})
|
|
|
|
|
(trustLevel := 1)
|
2024-01-16 14:11:52 -08:00
|
|
|
|
let testCases: List (String × ConstantCat) := [
|
|
|
|
|
("Or", ConstantCat.induct {
|
2024-01-16 13:29:30 -08:00
|
|
|
|
numParams := 2,
|
|
|
|
|
numIndices := 0,
|
2024-03-14 22:40:14 -07:00
|
|
|
|
all := #["Or"],
|
|
|
|
|
ctors := #["Or.inl", "Or.inr"],
|
2024-01-16 14:11:52 -08:00
|
|
|
|
}),
|
|
|
|
|
("Except.ok", ConstantCat.ctor {
|
|
|
|
|
induct := "Except",
|
|
|
|
|
cidx := 1,
|
|
|
|
|
numParams := 2,
|
|
|
|
|
numFields := 1,
|
|
|
|
|
}),
|
|
|
|
|
("Eq.rec", ConstantCat.recursor {
|
2024-03-14 22:40:14 -07:00
|
|
|
|
all := #["Eq"],
|
2024-01-16 14:11:52 -08:00
|
|
|
|
numParams := 2,
|
|
|
|
|
numIndices := 1,
|
|
|
|
|
numMotives := 1,
|
|
|
|
|
numMinors := 1,
|
2024-04-09 21:24:08 -07:00
|
|
|
|
rules := #[{ ctor := "Eq.refl", nFields := 0, rhs := { pp? := .some "fun {α} a motive refl => refl" } }]
|
2024-01-16 14:11:52 -08:00
|
|
|
|
k := true,
|
|
|
|
|
}),
|
|
|
|
|
("ForM.rec", ConstantCat.recursor {
|
2024-03-14 22:40:14 -07:00
|
|
|
|
all := #["ForM"],
|
2024-01-16 14:11:52 -08:00
|
|
|
|
numParams := 3,
|
|
|
|
|
numIndices := 0,
|
|
|
|
|
numMotives := 1,
|
|
|
|
|
numMinors := 1,
|
2024-04-09 21:24:08 -07:00
|
|
|
|
rules := #[{ ctor := "ForM.mk", nFields := 1, rhs := { pp? := .some "fun m γ α motive mk forM => mk forM" } }]
|
2024-01-16 14:11:52 -08:00
|
|
|
|
k := false,
|
2024-01-16 13:29:30 -08:00
|
|
|
|
})
|
2024-01-16 14:11:52 -08:00
|
|
|
|
]
|
|
|
|
|
let inner: CoreM LSpec.TestSeq := do
|
|
|
|
|
testCases.foldlM (λ acc (name, cat) => do
|
|
|
|
|
let args: Protocol.EnvInspect := { name := name }
|
|
|
|
|
let result ← match ← Environment.inspect args (options := {}) with
|
|
|
|
|
| .ok result => pure $ result
|
|
|
|
|
| .error e => panic! s!"Error: {e.desc}"
|
|
|
|
|
let p := match cat with
|
|
|
|
|
| .induct info => LSpec.test name (result.inductInfo? == .some info)
|
|
|
|
|
| .ctor info => LSpec.test name (result.constructorInfo? == .some info)
|
|
|
|
|
| .recursor info => LSpec.test name (result.recursorInfo? == .some info)
|
|
|
|
|
return LSpec.TestSeq.append acc p
|
|
|
|
|
) LSpec.TestSeq.done
|
2024-01-16 13:29:30 -08:00
|
|
|
|
runCoreMSeq env inner
|
|
|
|
|
|
2024-04-06 14:07:13 -07:00
|
|
|
|
def suite: List (String × IO LSpec.TestSeq) :=
|
|
|
|
|
[
|
|
|
|
|
("Catalog", test_catalog),
|
|
|
|
|
("Symbol Visibility", test_symbol_visibility),
|
|
|
|
|
("Inspect", test_inspect),
|
|
|
|
|
]
|
2023-12-05 19:07:00 -08:00
|
|
|
|
|
2023-12-26 09:22:57 -08:00
|
|
|
|
end Pantograph.Test.Environment
|