2023-05-22 14:49:56 -07:00
|
|
|
|
import LSpec
|
2023-05-24 00:54:48 -07:00
|
|
|
|
import Pantograph.Serial
|
2023-05-22 14:49:56 -07:00
|
|
|
|
import Pantograph.Symbols
|
|
|
|
|
|
|
|
|
|
namespace Pantograph.Test
|
|
|
|
|
|
2023-05-24 00:54:48 -07:00
|
|
|
|
open Pantograph
|
|
|
|
|
open Lean
|
|
|
|
|
|
2023-08-14 21:43:40 -07:00
|
|
|
|
deriving instance Repr, DecidableEq for Commands.BoundExpression
|
|
|
|
|
|
|
|
|
|
def test_str_to_name: LSpec.TestSeq :=
|
|
|
|
|
LSpec.test "Symbol parsing" (Name.str (.str (.str .anonymous "Lean") "Meta") "run" = Pantograph.str_to_name "Lean.Meta.run")
|
2023-05-24 00:54:48 -07:00
|
|
|
|
|
|
|
|
|
def test_expr_to_binder (env: Environment): IO LSpec.TestSeq := do
|
2023-08-14 21:43:40 -07:00
|
|
|
|
let entries: List (String × Commands.BoundExpression) := [
|
2023-05-24 00:54:48 -07:00
|
|
|
|
("Nat.add_comm", { binders := #[("n", "Nat"), ("m", "Nat")], target := "n + m = m + n" }),
|
|
|
|
|
("Nat.le_of_succ_le", { binders := #[("n", "Nat"), ("m", "Nat"), ("h", "Nat.succ n ≤ m")], target := "n ≤ m" })
|
|
|
|
|
]
|
2023-08-14 21:43:40 -07:00
|
|
|
|
let coreM := entries.foldlM (λ suites (symbol, target) => do
|
2023-05-24 00:54:48 -07:00
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let expr := str_to_name symbol |> env.find? |>.get! |>.type
|
|
|
|
|
let test := LSpec.check symbol ((← type_expr_to_bound expr) = target)
|
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done |>.run'
|
|
|
|
|
let coreContext: Core.Context := {
|
2023-08-14 21:43:40 -07:00
|
|
|
|
currNamespace := Lean.Name.str .anonymous "Aniva"
|
2023-05-24 00:54:48 -07:00
|
|
|
|
openDecls := [], -- No 'open' directives needed
|
2023-08-14 21:43:40 -07:00
|
|
|
|
fileName := "<Pantograph/Test>",
|
2023-05-24 00:54:48 -07:00
|
|
|
|
fileMap := { source := "", positions := #[0], lines := #[1] }
|
|
|
|
|
}
|
|
|
|
|
match ← (coreM.run' coreContext { env := env }).toBaseIO with
|
|
|
|
|
| .error exception =>
|
|
|
|
|
return LSpec.test "Exception" (s!"internal exception #{← exception.toMessageData.toString}" = "")
|
|
|
|
|
| .ok a => return a
|
|
|
|
|
|
2023-08-14 21:43:40 -07:00
|
|
|
|
def test_sexp_of_symbol (env: Environment): IO LSpec.TestSeq := do
|
|
|
|
|
let entries: List (String × String) := [
|
|
|
|
|
-- This one contains unhygienic variable names which must be suppressed
|
|
|
|
|
("Nat.add", "(:forall :anon (:c Nat) (:forall :anon (:c Nat) (:c Nat)))"),
|
|
|
|
|
-- These ones are normal and easy
|
|
|
|
|
("Nat.add_one", "(:forall n (:c Nat) ((((:c Eq) (:c Nat)) (((((((:c HAdd.hAdd) (:c Nat)) (:c Nat)) (:c Nat)) (((:c instHAdd) (:c Nat)) (:c instAddNat))) 0) ((((:c OfNat.ofNat) (:c Nat)) (:lit 1)) ((:c instOfNatNat) (:lit 1))))) ((:c Nat.succ) 0)))"),
|
2023-08-23 12:51:06 -07:00
|
|
|
|
("Nat.le_of_succ_le", "(:forall n (:c Nat) (:forall m (:c Nat) (:forall h (((((:c LE.le) (:c Nat)) (:c instLENat)) ((:c Nat.succ) 1)) 0) (((((:c LE.le) (:c Nat)) (:c instLENat)) 2) 1)) :implicit) :implicit)"),
|
|
|
|
|
-- Handling of higher order types
|
|
|
|
|
("Or", "(:forall a (:sort 0) (:forall b (:sort 0) (:sort 0)))"),
|
|
|
|
|
("List", "(:forall α (:sort (+ u 1)) (:sort (+ u 1)))")
|
2023-08-14 21:43:40 -07:00
|
|
|
|
]
|
|
|
|
|
let metaM: MetaM LSpec.TestSeq := entries.foldlM (λ suites (symbol, target) => do
|
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let expr := str_to_name symbol |> env.find? |>.get! |>.type
|
|
|
|
|
let test := LSpec.check symbol ((← serialize_expression_ast expr) = target)
|
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done |>.run'
|
|
|
|
|
let coreM := metaM.run'
|
|
|
|
|
let coreContext: Core.Context := {
|
|
|
|
|
currNamespace := Lean.Name.str .anonymous "Aniva"
|
|
|
|
|
openDecls := [], -- No 'open' directives needed
|
|
|
|
|
fileName := "<Pantograph/Test>",
|
|
|
|
|
fileMap := { source := "", positions := #[0], lines := #[1] }
|
|
|
|
|
}
|
|
|
|
|
match ← (coreM.run' coreContext { env := env }).toBaseIO with
|
|
|
|
|
| .error exception =>
|
|
|
|
|
return LSpec.test "Exception" (s!"internal exception #{← exception.toMessageData.toString}" = "")
|
|
|
|
|
| .ok a => return a
|
|
|
|
|
|
|
|
|
|
|
2023-05-22 14:49:56 -07:00
|
|
|
|
def test_serial: IO LSpec.TestSeq := do
|
2023-05-24 00:54:48 -07:00
|
|
|
|
let env: Environment ← importModules
|
|
|
|
|
(imports := ["Init"].map (λ str => { module := str_to_name str, runtimeOnly := false }))
|
|
|
|
|
(opts := {})
|
|
|
|
|
(trustLevel := 1)
|
|
|
|
|
|
2023-05-22 14:49:56 -07:00
|
|
|
|
return LSpec.group "Serialisation" $
|
2023-08-14 21:43:40 -07:00
|
|
|
|
(LSpec.group "str_to_name" test_str_to_name) ++
|
2023-05-24 00:54:48 -07:00
|
|
|
|
(LSpec.group "Expression binder" (← test_expr_to_binder env)) ++
|
2023-08-14 21:43:40 -07:00
|
|
|
|
(LSpec.group "Sexp from symbol" (← test_sexp_of_symbol env))
|
2023-05-22 14:49:56 -07:00
|
|
|
|
|
|
|
|
|
end Pantograph.Test
|