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
|
|
|
|
|
2024-04-12 12:37:37 -07:00
|
|
|
|
def test_serializeName: LSpec.TestSeq :=
|
2023-10-30 14:45:43 -07:00
|
|
|
|
let quote := "\""
|
|
|
|
|
let escape := "\\"
|
2024-04-12 12:37:37 -07:00
|
|
|
|
LSpec.test "a.b.1" (serializeName (Name.num (.str (.str .anonymous "a") "b") 1) = "a.b.1") ++
|
|
|
|
|
LSpec.test "seg.«a.b»" (serializeName (Name.str (.str .anonymous "seg") "a.b") = s!"{quote}seg.«a.b»{quote}") ++
|
2023-10-29 11:56:56 -07:00
|
|
|
|
-- Pathological test case
|
2024-04-12 12:37:37 -07:00
|
|
|
|
LSpec.test s!"«̈{escape}{quote}»" (serializeName (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
|
|
|
|
]
|
2024-04-09 10:03:36 -07:00
|
|
|
|
runCoreMSeq env $ 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
|
2024-04-12 16:34:21 -07:00
|
|
|
|
let test := LSpec.check symbol.toString ((← typeExprToBound expr) = target)
|
2023-05-24 00:54:48 -07:00
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done |>.run'
|
|
|
|
|
|
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
|
|
|
|
]
|
2024-04-09 10:03:36 -07:00
|
|
|
|
runMetaMSeq env $ entries.foldlM (λ suites (symbol, target) => do
|
2023-08-14 21:43:40 -07:00
|
|
|
|
let env ← MonadEnv.getEnv
|
2023-11-06 10:45:11 -08:00
|
|
|
|
let expr := env.find? symbol.toName |>.get! |>.type
|
2024-04-12 12:37:37 -07:00
|
|
|
|
let test := LSpec.check symbol ((← serializeExpressionSexp expr) = target)
|
2023-10-25 22:18:59 -07:00
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2024-04-09 10:03:36 -07:00
|
|
|
|
def test_sexp_of_elab (env: Environment): IO LSpec.TestSeq := do
|
2024-06-11 12:44:42 -07:00
|
|
|
|
let entries: List (String × (List Name) × 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))"),
|
|
|
|
|
("λ {α: Sort (u + 1)} => List α", [`u], "(:lambda α (:sort (+ u 1)) ((:c List) 0) :implicit)"),
|
|
|
|
|
("λ {α} => List α", [], "(:lambda α (:sort (+ (:mv _uniq.4) 1)) ((:c List) 0) :implicit)"),
|
2024-06-11 15:21:35 -07:00
|
|
|
|
("(2: Nat) <= (5: Nat)", [], "((:c LE.le) (:mv _uniq.18) (:mv _uniq.19) ((:c OfNat.ofNat) (:mv _uniq.4) (:lit 2) (:mv _uniq.5)) ((:c OfNat.ofNat) (:mv _uniq.14) (:lit 5) (:mv _uniq.15)))"),
|
2023-12-04 16:21:02 -08:00
|
|
|
|
]
|
2024-06-11 12:44:42 -07:00
|
|
|
|
entries.foldlM (λ suites (source, levels, target) =>
|
|
|
|
|
let termElabM := do
|
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let s ← match parseTerm env source with
|
|
|
|
|
| .ok s => pure s
|
|
|
|
|
| .error e => return parseFailure e
|
|
|
|
|
let expr ← match (← elabTerm s) with
|
|
|
|
|
| .ok expr => pure expr
|
|
|
|
|
| .error e => return elabFailure e
|
|
|
|
|
return LSpec.check source ((← serializeExpressionSexp expr) = target)
|
2024-07-22 17:57:01 -07:00
|
|
|
|
let metaM := (Elab.Term.withLevelNames levels termElabM).run' (ctx := Condensed.elabContext)
|
2024-06-11 12:44:42 -07:00
|
|
|
|
return LSpec.TestSeq.append suites (← runMetaMSeq env metaM))
|
|
|
|
|
LSpec.TestSeq.done
|
2024-04-09 10:03:36 -07:00
|
|
|
|
|
|
|
|
|
def test_sexp_of_expr (env: Environment): IO LSpec.TestSeq := do
|
|
|
|
|
let entries: List (Expr × String) := [
|
|
|
|
|
(.lam `p (.sort .zero)
|
|
|
|
|
(.lam `q (.sort .zero)
|
|
|
|
|
(.lam `k (mkApp2 (.const `And []) (.bvar 1) (.bvar 0))
|
|
|
|
|
(.proj `And 1 (.bvar 0))
|
|
|
|
|
.default)
|
|
|
|
|
.implicit)
|
|
|
|
|
.implicit,
|
2024-04-09 10:06:26 -07:00
|
|
|
|
"(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda k ((:c And) 1 0) ((:c And.right) _ _ 0)) :implicit) :implicit)"
|
2024-04-09 10:03:36 -07:00
|
|
|
|
),
|
|
|
|
|
]
|
|
|
|
|
let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (expr, target) => do
|
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let testCaseName := target.take 10
|
2024-04-12 12:37:37 -07:00
|
|
|
|
let test := LSpec.check testCaseName ((← serializeExpressionSexp expr) = target)
|
2024-04-09 10:03:36 -07:00
|
|
|
|
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done
|
2024-07-22 17:57:01 -07:00
|
|
|
|
runMetaMSeq env $ termElabM.run' (ctx := Condensed.elabContext)
|
2024-01-07 14:14:20 -08:00
|
|
|
|
|
|
|
|
|
-- Instance parsing
|
2024-04-09 10:03:36 -07:00
|
|
|
|
def test_instance (env: Environment): IO LSpec.TestSeq :=
|
|
|
|
|
runMetaMSeq env do
|
2024-01-07 14:14:20 -08:00
|
|
|
|
let env ← MonadEnv.getEnv
|
|
|
|
|
let source := "λ x y: Nat => HAdd.hAdd Nat Nat Nat (instHAdd Nat instAddNat) x y"
|
2024-03-31 15:55:08 -07:00
|
|
|
|
let s := parseTerm env source |>.toOption |>.get!
|
|
|
|
|
let _expr := (← runTermElabMInMeta <| elabTerm s) |>.toOption |>.get!
|
2024-01-07 14:14:20 -08:00
|
|
|
|
return LSpec.TestSeq.done
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2024-04-06 14:07:13 -07:00
|
|
|
|
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
|
|
|
|
[
|
2024-04-12 12:37:37 -07:00
|
|
|
|
("serializeName", do pure test_serializeName),
|
2024-04-06 14:07:13 -07:00
|
|
|
|
("Expression binder", test_expr_to_binder env),
|
|
|
|
|
("Sexp from symbol", test_sexp_of_symbol env),
|
2024-04-09 10:03:36 -07:00
|
|
|
|
("Sexp from elaborated expr", test_sexp_of_elab env),
|
2024-04-06 14:07:13 -07:00
|
|
|
|
("Sexp from expr", test_sexp_of_expr env),
|
|
|
|
|
("Instance", test_instance env),
|
|
|
|
|
]
|
2023-05-22 14:49:56 -07:00
|
|
|
|
|
2023-10-06 17:31:36 -07:00
|
|
|
|
end Pantograph.Test.Serial
|