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
|
2023-12-05 19:07:00 -08:00
|
|
|
|
|
2023-12-26 09:22:57 -08:00
|
|
|
|
namespace Pantograph.Test.Environment
|
2023-12-05 19:07:00 -08:00
|
|
|
|
|
|
|
|
|
open Pantograph
|
|
|
|
|
open Lean
|
|
|
|
|
|
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
|
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.RecursorInfo
|
2024-01-16 13:29:30 -08:00
|
|
|
|
deriving instance DecidableEq, Repr for Protocol.EnvInspectResult
|
|
|
|
|
|
2023-12-05 19:07:00 -08:00
|
|
|
|
def test_symbol_visibility (env: Environment): IO LSpec.TestSeq := do
|
|
|
|
|
let entries: List (Name × Bool) := [
|
|
|
|
|
("Nat.add_comm".toName, false),
|
2023-12-08 16:17:16 -08:00
|
|
|
|
("Lean.Name".toName, true)
|
2023-12-05 19:07:00 -08:00
|
|
|
|
]
|
|
|
|
|
let suite := entries.foldl (λ suites (symbol, target) =>
|
|
|
|
|
let constant := env.constants.find! symbol
|
2023-12-15 10:40:36 -08:00
|
|
|
|
let test := LSpec.check symbol.toString ((Environment.is_symbol_unsafe_or_internal symbol constant) == 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-01-16 13:29:30 -08:00
|
|
|
|
def test_inspect (env: Environment): IO LSpec.TestSeq := do
|
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,
|
|
|
|
|
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 {
|
|
|
|
|
all := ["Eq"],
|
|
|
|
|
numParams := 2,
|
|
|
|
|
numIndices := 1,
|
|
|
|
|
numMotives := 1,
|
|
|
|
|
numMinors := 1,
|
|
|
|
|
k := true,
|
|
|
|
|
}),
|
|
|
|
|
("ForM.rec", ConstantCat.recursor {
|
|
|
|
|
all := ["ForM"],
|
|
|
|
|
numParams := 3,
|
|
|
|
|
numIndices := 0,
|
|
|
|
|
numMotives := 1,
|
|
|
|
|
numMinors := 1,
|
|
|
|
|
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
|
|
|
|
|
|
2023-12-05 19:07:00 -08:00
|
|
|
|
def suite: IO LSpec.TestSeq := do
|
|
|
|
|
let env: Environment ← importModules
|
|
|
|
|
(imports := #["Init"].map (λ str => { module := str.toName, runtimeOnly := false }))
|
|
|
|
|
(opts := {})
|
|
|
|
|
(trustLevel := 1)
|
|
|
|
|
|
2023-12-26 09:22:57 -08:00
|
|
|
|
return LSpec.group "Environment" $
|
2024-01-16 13:29:30 -08:00
|
|
|
|
(LSpec.group "Symbol visibility" (← test_symbol_visibility env)) ++
|
|
|
|
|
(LSpec.group "Inspect" (← test_inspect env))
|
2023-12-05 19:07:00 -08:00
|
|
|
|
|
2023-12-26 09:22:57 -08:00
|
|
|
|
end Pantograph.Test.Environment
|