2023-05-22 14:49:56 -07:00
|
|
|
|
import LSpec
|
2023-05-24 00:54:48 -07:00
|
|
|
|
import Pantograph.Serial
|
2024-01-07 14:14:20 -08:00
|
|
|
|
import Test.Common
|
2024-03-28 18:56:42 -07:00
|
|
|
|
import Lean
|
2023-05-22 14:49:56 -07:00
|
|
|
|
|
2024-03-28 18:56:42 -07:00
|
|
|
|
open Lean
|
2023-10-06 17:31:36 -07:00
|
|
|
|
namespace Pantograph.Test.Serial
|
2023-05-22 14:49:56 -07:00
|
|
|
|
|
2023-05-24 00:54:48 -07:00
|
|
|
|
open Pantograph
|
|
|
|
|
|
2023-10-15 17:15:23 -07:00
|
|
|
|
deriving instance Repr, DecidableEq for Protocol.BoundExpression
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2023-10-29 11:56:56 -07:00
|
|
|
|
def test_name_to_ast: LSpec.TestSeq :=
|
2023-10-30 14:45:43 -07:00
|
|
|
|
let quote := "\""
|
|
|
|
|
let escape := "\\"
|
2023-10-29 11:56:56 -07:00
|
|
|
|
LSpec.test "a.b.1" (name_to_ast (Name.num (.str (.str .anonymous "a") "b") 1) = "a.b.1") ++
|
2023-10-30 14:45:43 -07:00
|
|
|
|
LSpec.test "seg.«a.b»" (name_to_ast (Name.str (.str .anonymous "seg") "a.b") = s!"{quote}seg.«a.b»{quote}") ++
|
2023-10-29 11:56:56 -07:00
|
|
|
|
-- Pathological test case
|
2023-10-30 14:45:43 -07:00
|
|
|
|
LSpec.test s!"«̈{escape}{quote}»" (name_to_ast (Name.str .anonymous s!"{escape}{quote}") = s!"{quote}«{escape}{quote}»{quote}")
|
2023-05-24 00:54:48 -07:00
|
|
|
|
|
|
|
|
|
def test_expr_to_binder (env: Environment): IO LSpec.TestSeq := do
|
2023-11-06 10:45:11 -08:00
|
|
|
|
let entries: List (Name × Protocol.BoundExpression) := [
|
|
|
|
|
("Nat.add_comm".toName, { binders := #[("n", "Nat"), ("m", "Nat")], target := "n + m = m + n" }),
|
2024-03-28 00:06:35 -07:00
|
|
|
|
("Nat.le_of_succ_le".toName, { binders := #[("n", "Nat"), ("m", "Nat"), ("h", "n.succ ≤ m")], target := "n ≤ m" })
|
2023-05-24 00:54:48 -07:00
|
|
|
|
]
|
2023-11-06 10:45:11 -08:00
|
|
|
|
let coreM: CoreM LSpec.TestSeq := entries.foldlM (λ suites (symbol, target) => do
|
2023-05-24 00:54:48 -07:00
|
|
|
|
let env ← MonadEnv.getEnv
|
2023-11-06 10:45:11 -08:00
|
|
|
|
let expr := env.find? symbol |>.get! |>.type
|
|
|
|
|
let test := LSpec.check symbol.toString ((← type_expr_to_bound expr) = target)
|
2023-05-24 00:54:48 -07:00
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done |>.run'
|
2024-01-07 14:14:20 -08:00
|
|
|
|
runCoreMSeq env coreM
|
2023-05-24 00:54:48 -07:00
|
|
|
|
|
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
|
2023-10-29 11:18:35 -07:00
|
|
|
|
("Nat.add", "(:forall _ (:c Nat) (:forall _ (:c Nat) (:c Nat)))"),
|
2023-08-14 21:43:40 -07:00
|
|
|
|
-- These ones are normal and easy
|
2023-10-29 12:50:36 -07:00
|
|
|
|
("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)))"),
|
|
|
|
|
("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)"),
|
2023-08-23 12:51:06 -07:00
|
|
|
|
-- 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
|
2023-11-06 10:45:11 -08:00
|
|
|
|
let expr := env.find? symbol.toName |>.get! |>.type
|
2023-12-04 16:21:02 -08:00
|
|
|
|
let test := LSpec.check symbol ((← serialize_expression_ast expr) = target)
|
2023-10-25 22:18:59 -07:00
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done
|
2024-01-07 14:14:20 -08:00
|
|
|
|
runMetaMSeq env metaM
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2023-12-04 16:21:02 -08:00
|
|
|
|
def test_sexp_of_expr (env: Environment): IO LSpec.TestSeq := do
|
|
|
|
|
let entries: List (String × String) := [
|
|
|
|
|
("λ x: Nat × Bool => x.1", "(:lambda x ((:c Prod) (:c Nat) (:c Bool)) ((:c Prod.fst) (:c Nat) (:c Bool) 0))"),
|
|
|
|
|
("λ x: Array Nat => x.data", "(:lambda x ((:c Array) (:c Nat)) ((:c Array.data) (:c Nat) 0))")
|
|
|
|
|
]
|
|
|
|
|
let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (source, target) => do
|
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let s := syntax_from_str env source |>.toOption |>.get!
|
|
|
|
|
let expr := (← syntax_to_expr s) |>.toOption |>.get!
|
|
|
|
|
let test := LSpec.check source ((← serialize_expression_ast expr) = target)
|
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done
|
2024-03-28 00:06:35 -07:00
|
|
|
|
let metaM := termElabM.run' (ctx := defaultTermElabMContext)
|
2024-01-07 14:14:20 -08:00
|
|
|
|
runMetaMSeq env metaM
|
|
|
|
|
|
|
|
|
|
-- Instance parsing
|
|
|
|
|
def test_instance (env: Environment): IO LSpec.TestSeq := do
|
|
|
|
|
let metaM: MetaM LSpec.TestSeq := do
|
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let source := "λ x y: Nat => HAdd.hAdd Nat Nat Nat (instHAdd Nat instAddNat) x y"
|
|
|
|
|
let s := syntax_from_str env source |>.toOption |>.get!
|
2024-01-16 14:11:52 -08:00
|
|
|
|
let _expr := (← runTermElabMInMeta <| syntax_to_expr s) |>.toOption |>.get!
|
2024-01-07 14:14:20 -08:00
|
|
|
|
return LSpec.TestSeq.done
|
|
|
|
|
runMetaMSeq env metaM
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2023-10-06 17:31:36 -07:00
|
|
|
|
def suite: IO LSpec.TestSeq := do
|
2023-05-24 00:54:48 -07:00
|
|
|
|
let env: Environment ← importModules
|
2023-11-06 10:45:11 -08:00
|
|
|
|
(imports := #["Init"].map (λ str => { module := str.toName, runtimeOnly := false }))
|
2023-05-24 00:54:48 -07:00
|
|
|
|
(opts := {})
|
|
|
|
|
(trustLevel := 1)
|
|
|
|
|
|
2023-10-06 17:31:36 -07:00
|
|
|
|
return LSpec.group "Serialization" $
|
2023-10-29 11:56:56 -07:00
|
|
|
|
(LSpec.group "name_to_ast" test_name_to_ast) ++
|
2023-05-24 00:54:48 -07:00
|
|
|
|
(LSpec.group "Expression binder" (← test_expr_to_binder env)) ++
|
2023-12-04 16:21:02 -08:00
|
|
|
|
(LSpec.group "Sexp from symbol" (← test_sexp_of_symbol env)) ++
|
2024-01-07 14:14:20 -08:00
|
|
|
|
(LSpec.group "Sexp from expr" (← test_sexp_of_expr env)) ++
|
|
|
|
|
(LSpec.group "Instance" (← test_instance env))
|
2023-05-22 14:49:56 -07:00
|
|
|
|
|
2023-10-06 17:31:36 -07:00
|
|
|
|
end Pantograph.Test.Serial
|