2023-08-14 21:43:40 -07:00
|
|
|
|
/- Integration test for the REPL
|
|
|
|
|
-/
|
|
|
|
|
import LSpec
|
|
|
|
|
import Pantograph
|
2024-09-06 22:01:36 -07:00
|
|
|
|
import Repl
|
2024-09-08 11:57:39 -07:00
|
|
|
|
import Test.Common
|
2024-09-06 22:01:36 -07:00
|
|
|
|
|
2023-10-06 17:31:36 -07:00
|
|
|
|
namespace Pantograph.Test.Integration
|
2024-09-08 15:02:43 -07:00
|
|
|
|
open Pantograph.Repl
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def step { α } [Lean.ToJson α] (cmd: String) (payload: List (String × Lean.Json))
|
|
|
|
|
(expected: α) (name? : Option String := .none): MainM LSpec.TestSeq := do
|
|
|
|
|
let payload := Lean.Json.mkObj payload
|
|
|
|
|
let name := name?.getD s!"{cmd} {payload.compress}"
|
2024-09-08 15:02:43 -07:00
|
|
|
|
let result ← Repl.execute { cmd, payload }
|
2024-09-08 11:57:39 -07:00
|
|
|
|
return LSpec.test name (toString result = toString (Lean.toJson expected))
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2024-09-08 11:57:39 -07:00
|
|
|
|
abbrev Test := List (MainM LSpec.TestSeq)
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def test_elab : Test :=
|
|
|
|
|
[
|
|
|
|
|
step "expr.echo"
|
2024-06-11 12:44:42 -07:00
|
|
|
|
[("expr", .str "λ {α : Sort (u + 1)} => List α"), ("levels", .arr #["u"])]
|
2024-04-06 17:45:36 -07:00
|
|
|
|
(Lean.toJson ({
|
|
|
|
|
type := { pp? := .some "{α : Type u} → Type u" },
|
|
|
|
|
expr := { pp? := .some "fun {α} => List α" }
|
|
|
|
|
}: Protocol.ExprEchoResult)),
|
|
|
|
|
]
|
|
|
|
|
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def test_option_modify : Test :=
|
2024-03-28 00:06:35 -07:00
|
|
|
|
let pp? := Option.some "∀ (n : Nat), n + 1 = n.succ"
|
2023-10-29 12:50:36 -07:00
|
|
|
|
let sexp? := Option.some "(: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-14 21:43:40 -07:00
|
|
|
|
let module? := Option.some "Init.Data.Nat.Basic"
|
2023-10-15 17:15:23 -07:00
|
|
|
|
let options: Protocol.Options := {}
|
2024-09-08 11:57:39 -07:00
|
|
|
|
[
|
|
|
|
|
step "env.inspect" [("name", .str "Nat.add_one")]
|
|
|
|
|
({ type := { pp? }, module? }: Protocol.EnvInspectResult),
|
|
|
|
|
step "options.set" [("printExprAST", .bool true)]
|
|
|
|
|
({ }: Protocol.OptionsSetResult),
|
|
|
|
|
step "env.inspect" [("name", .str "Nat.add_one")]
|
|
|
|
|
({ type := { pp?, sexp? }, module? }: Protocol.EnvInspectResult),
|
|
|
|
|
step "options.print" []
|
|
|
|
|
({ options with printExprAST := true }: Protocol.Options),
|
2023-08-14 21:43:40 -07:00
|
|
|
|
]
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def test_malformed_command : Test :=
|
2023-08-23 13:00:11 -07:00
|
|
|
|
let invalid := "invalid"
|
2024-09-08 11:57:39 -07:00
|
|
|
|
[
|
|
|
|
|
step invalid [("name", .str "Nat.add_one")]
|
|
|
|
|
({ error := "command", desc := s!"Unknown command {invalid}" }: Protocol.InteractionError)
|
|
|
|
|
(name? := .some "Invalid Command"),
|
|
|
|
|
step "expr.echo" [(invalid, .str "Random garbage data")]
|
|
|
|
|
({ error := "command", desc := s!"Unable to parse json: Pantograph.Protocol.ExprEcho.expr: String expected" }:
|
|
|
|
|
Protocol.InteractionError)
|
|
|
|
|
(name? := .some "JSON Deserialization")
|
2023-08-23 13:00:11 -07:00
|
|
|
|
]
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def test_tactic : Test :=
|
2023-11-07 13:07:50 -08:00
|
|
|
|
let goal1: Protocol.Goal := {
|
2024-03-31 16:06:30 -07:00
|
|
|
|
name := "_uniq.11",
|
2023-10-27 15:32:59 -07:00
|
|
|
|
target := { pp? := .some "∀ (q : Prop), x ∨ q → q ∨ x" },
|
2024-08-17 00:53:38 -07:00
|
|
|
|
vars := #[{ name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }}],
|
2023-10-27 15:32:59 -07:00
|
|
|
|
}
|
2023-11-07 13:07:50 -08:00
|
|
|
|
let goal2: Protocol.Goal := {
|
2024-05-15 21:34:15 -07:00
|
|
|
|
name := "_uniq.17",
|
2023-11-07 13:07:50 -08:00
|
|
|
|
target := { pp? := .some "x ∨ y → y ∨ x" },
|
|
|
|
|
vars := #[
|
2024-08-17 00:53:38 -07:00
|
|
|
|
{ name := "_uniq.10", userName := "x", type? := .some { pp? := .some "Prop" }},
|
|
|
|
|
{ name := "_uniq.16", userName := "y", type? := .some { pp? := .some "Prop" }}
|
2023-11-07 13:07:50 -08:00
|
|
|
|
],
|
|
|
|
|
}
|
2024-09-08 11:57:39 -07:00
|
|
|
|
[
|
|
|
|
|
step "goal.start" [("expr", .str "∀ (p q: Prop), p ∨ q → q ∨ p")]
|
|
|
|
|
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult),
|
|
|
|
|
step "goal.tactic" [("stateId", .num 0), ("goalId", .num 0), ("tactic", .str "intro x")]
|
|
|
|
|
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
|
|
|
|
|
step "goal.print" [("stateId", .num 1)]
|
|
|
|
|
({ parent? := .some { pp? := .some "fun x => ?m.12 x" }, }: Protocol.GoalPrintResult),
|
|
|
|
|
step "goal.tactic" [("stateId", .num 1), ("goalId", .num 0), ("tactic", .str "intro y")]
|
|
|
|
|
({ nextStateId? := .some 2, goals? := #[goal2], }: Protocol.GoalTacticResult),
|
|
|
|
|
]
|
|
|
|
|
def test_automatic_mode (automatic: Bool): Test :=
|
|
|
|
|
let varsPQ := #[
|
|
|
|
|
{ name := "_uniq.10", userName := "p", type? := .some { pp? := .some "Prop" }},
|
|
|
|
|
{ name := "_uniq.13", userName := "q", type? := .some { pp? := .some "Prop" }}
|
|
|
|
|
]
|
|
|
|
|
let goal1: Protocol.Goal := {
|
|
|
|
|
name := "_uniq.17",
|
|
|
|
|
target := { pp? := .some "q ∨ p" },
|
|
|
|
|
vars := varsPQ ++ #[
|
|
|
|
|
{ name := "_uniq.16", userName := "h", type? := .some { pp? := .some "p ∨ q" }}
|
|
|
|
|
],
|
|
|
|
|
}
|
|
|
|
|
let goal2l: Protocol.Goal := {
|
|
|
|
|
name := "_uniq.59",
|
|
|
|
|
userName? := .some "inl",
|
|
|
|
|
target := { pp? := .some "q ∨ p" },
|
|
|
|
|
vars := varsPQ ++ #[
|
|
|
|
|
{ name := "_uniq.47", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true}
|
|
|
|
|
],
|
|
|
|
|
}
|
|
|
|
|
let goal2r: Protocol.Goal := {
|
|
|
|
|
name := "_uniq.72",
|
|
|
|
|
userName? := .some "inr",
|
|
|
|
|
target := { pp? := .some "q ∨ p" },
|
|
|
|
|
vars := varsPQ ++ #[
|
|
|
|
|
{ name := "_uniq.60", userName := "h✝", type? := .some { pp? := .some "q" }, isInaccessible := true}
|
|
|
|
|
],
|
|
|
|
|
}
|
|
|
|
|
let goal3l: Protocol.Goal := {
|
|
|
|
|
name := "_uniq.78",
|
|
|
|
|
userName? := .some "inl.h",
|
|
|
|
|
target := { pp? := .some "p" },
|
|
|
|
|
vars := varsPQ ++ #[
|
|
|
|
|
{ name := "_uniq.47", userName := "h✝", type? := .some { pp? := .some "p" }, isInaccessible := true}
|
|
|
|
|
],
|
|
|
|
|
}
|
|
|
|
|
[
|
|
|
|
|
step "options.set" [("automaticMode", .bool automatic)]
|
|
|
|
|
({}: Protocol.OptionsSetResult),
|
|
|
|
|
step "goal.start" [("expr", .str "∀ (p q: Prop), p ∨ q → q ∨ p")]
|
|
|
|
|
({ stateId := 0, root := "_uniq.9" }: Protocol.GoalStartResult),
|
|
|
|
|
step "goal.tactic" [("stateId", .num 0), ("goalId", .num 0), ("tactic", .str "intro p q h")]
|
|
|
|
|
({ nextStateId? := .some 1, goals? := #[goal1], }: Protocol.GoalTacticResult),
|
|
|
|
|
step "goal.tactic" [("stateId", .num 1), ("goalId", .num 0), ("tactic", .str "cases h")]
|
|
|
|
|
({ nextStateId? := .some 2, goals? := #[goal2l, goal2r], }: Protocol.GoalTacticResult),
|
|
|
|
|
let goals? := if automatic then #[goal3l, goal2r] else #[goal3l]
|
|
|
|
|
step "goal.tactic" [("stateId", .num 2), ("goalId", .num 0), ("tactic", .str "apply Or.inr")]
|
|
|
|
|
({ nextStateId? := .some 3, goals?, }: Protocol.GoalTacticResult),
|
2023-10-27 15:32:59 -07:00
|
|
|
|
]
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def test_env_add_inspect : Test :=
|
2023-12-15 10:07:59 -08:00
|
|
|
|
let name1 := "Pantograph.mystery"
|
|
|
|
|
let name2 := "Pantograph.mystery2"
|
2024-09-08 11:57:39 -07:00
|
|
|
|
[
|
|
|
|
|
step "env.add"
|
2023-12-14 11:11:24 -08:00
|
|
|
|
[
|
2023-12-15 10:07:59 -08:00
|
|
|
|
("name", .str name1),
|
2023-12-14 11:11:24 -08:00
|
|
|
|
("type", .str "Prop → Prop → Prop"),
|
|
|
|
|
("value", .str "λ (a b: Prop) => Or a b"),
|
|
|
|
|
("isTheorem", .bool false)
|
|
|
|
|
]
|
2024-09-08 11:57:39 -07:00
|
|
|
|
({}: Protocol.EnvAddResult),
|
|
|
|
|
step "env.inspect" [("name", .str name1)]
|
|
|
|
|
({
|
2023-12-14 11:11:24 -08:00
|
|
|
|
value? := .some { pp? := .some "fun a b => a ∨ b" },
|
|
|
|
|
type := { pp? := .some "Prop → Prop → Prop" },
|
2023-12-15 10:07:59 -08:00
|
|
|
|
}:
|
2024-09-08 11:57:39 -07:00
|
|
|
|
Protocol.EnvInspectResult),
|
|
|
|
|
step "env.add"
|
2023-12-15 10:07:59 -08:00
|
|
|
|
[
|
|
|
|
|
("name", .str name2),
|
|
|
|
|
("type", .str "Nat → Int"),
|
|
|
|
|
("value", .str "λ (a: Nat) => a + 1"),
|
|
|
|
|
("isTheorem", .bool false)
|
|
|
|
|
]
|
2024-09-08 11:57:39 -07:00
|
|
|
|
({}: Protocol.EnvAddResult),
|
|
|
|
|
step "env.inspect" [("name", .str name2)]
|
|
|
|
|
({
|
2024-03-28 00:06:35 -07:00
|
|
|
|
value? := .some { pp? := .some "fun a => ↑a + 1" },
|
2023-12-15 10:07:59 -08:00
|
|
|
|
type := { pp? := .some "Nat → Int" },
|
2023-12-14 11:11:24 -08:00
|
|
|
|
}:
|
2024-09-08 11:57:39 -07:00
|
|
|
|
Protocol.EnvInspectResult)
|
2023-12-14 11:11:24 -08:00
|
|
|
|
]
|
|
|
|
|
|
2024-09-09 10:18:20 -07:00
|
|
|
|
example : ∀ (p: Prop), p → p := by
|
|
|
|
|
intro p h
|
|
|
|
|
exact h
|
|
|
|
|
|
|
|
|
|
def test_frontend_process : Test :=
|
|
|
|
|
[
|
|
|
|
|
let file := "example : ∀ (p: Prop), p → p := by\n intro p h\n exact h"
|
|
|
|
|
let goal1 := "p : Prop\nh : p\n⊢ p"
|
|
|
|
|
step "frontend.process"
|
|
|
|
|
[
|
|
|
|
|
("file", .str file),
|
|
|
|
|
("invocations", .bool true),
|
|
|
|
|
("sorrys", .bool false),
|
|
|
|
|
]
|
|
|
|
|
({
|
2024-10-05 14:49:17 -07:00
|
|
|
|
units := [{
|
|
|
|
|
boundary := (0, file.utf8ByteSize),
|
|
|
|
|
invocations? := .some [
|
|
|
|
|
{
|
|
|
|
|
goalBefore := "⊢ ∀ (p : Prop), p → p",
|
|
|
|
|
goalAfter := goal1,
|
|
|
|
|
tactic := "intro p h",
|
|
|
|
|
},
|
|
|
|
|
{
|
|
|
|
|
goalBefore := goal1 ,
|
|
|
|
|
goalAfter := "",
|
|
|
|
|
tactic := "exact h",
|
|
|
|
|
},
|
|
|
|
|
]
|
|
|
|
|
}],
|
2024-09-09 10:18:20 -07:00
|
|
|
|
}: Protocol.FrontendProcessResult),
|
|
|
|
|
]
|
|
|
|
|
|
2024-09-09 12:26:46 -07:00
|
|
|
|
example : 1 + 2 = 3 := rfl
|
|
|
|
|
example (p: Prop): p → p := by simp
|
|
|
|
|
|
|
|
|
|
def test_frontend_process_sorry : Test :=
|
|
|
|
|
let solved := "example : 1 + 2 = 3 := rfl\n"
|
|
|
|
|
let withSorry := "example (p: Prop): p → p := sorry"
|
|
|
|
|
[
|
|
|
|
|
let file := s!"{solved}{withSorry}"
|
|
|
|
|
let goal1: Protocol.Goal := {
|
2024-10-03 01:29:46 -07:00
|
|
|
|
name := "_uniq.6",
|
2024-09-09 12:26:46 -07:00
|
|
|
|
target := { pp? := .some "p → p" },
|
2024-10-03 01:29:46 -07:00
|
|
|
|
vars := #[{ name := "_uniq.4", userName := "p", type? := .some { pp? := .some "Prop" }}],
|
2024-09-09 12:26:46 -07:00
|
|
|
|
}
|
|
|
|
|
step "frontend.process"
|
|
|
|
|
[
|
|
|
|
|
("file", .str file),
|
|
|
|
|
("invocations", .bool false),
|
|
|
|
|
("sorrys", .bool true),
|
|
|
|
|
]
|
|
|
|
|
({
|
2024-10-05 14:49:17 -07:00
|
|
|
|
units := [{
|
|
|
|
|
boundary := (0, solved.utf8ByteSize),
|
|
|
|
|
}, {
|
|
|
|
|
boundary := (solved.utf8ByteSize, solved.utf8ByteSize + withSorry.utf8ByteSize),
|
|
|
|
|
goalStateId? := .some 0,
|
|
|
|
|
goals := #[goal1],
|
|
|
|
|
messages := #["<anonymous>:2:0: warning: declaration uses 'sorry'\n"],
|
|
|
|
|
}],
|
2024-09-09 12:26:46 -07:00
|
|
|
|
}: Protocol.FrontendProcessResult),
|
|
|
|
|
]
|
|
|
|
|
|
2024-09-09 10:18:20 -07:00
|
|
|
|
|
2024-09-08 11:57:39 -07:00
|
|
|
|
def runTest (env: Lean.Environment) (steps: Test): IO LSpec.TestSeq := do
|
|
|
|
|
-- Setup the environment for execution
|
|
|
|
|
let context: Context := {
|
|
|
|
|
imports := ["Init"]
|
|
|
|
|
}
|
|
|
|
|
let commands: MainM LSpec.TestSeq :=
|
|
|
|
|
steps.foldlM (λ suite step => do
|
|
|
|
|
let result ← step
|
|
|
|
|
return suite ++ result) LSpec.TestSeq.done
|
|
|
|
|
runCoreMSeq env <| commands.run context |>.run' {}
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
|
|
|
|
|
let tests := [
|
|
|
|
|
("expr.echo", test_elab),
|
|
|
|
|
("options.set options.print", test_option_modify),
|
2024-04-06 14:07:13 -07:00
|
|
|
|
("Malformed command", test_malformed_command),
|
|
|
|
|
("Tactic", test_tactic),
|
2024-09-08 11:57:39 -07:00
|
|
|
|
("Manual Mode", test_automatic_mode false),
|
|
|
|
|
("Automatic Mode", test_automatic_mode true),
|
2024-04-06 14:07:13 -07:00
|
|
|
|
("env.add env.inspect", test_env_add_inspect),
|
2024-09-09 12:26:46 -07:00
|
|
|
|
("frontend.process invocation", test_frontend_process),
|
|
|
|
|
("frontend.process sorry", test_frontend_process_sorry),
|
2024-04-06 14:07:13 -07:00
|
|
|
|
]
|
2024-09-08 11:57:39 -07:00
|
|
|
|
tests.map (fun (name, test) => (name, runTest env test))
|
2023-08-14 21:43:40 -07:00
|
|
|
|
|
|
|
|
|
|
2023-10-06 17:31:36 -07:00
|
|
|
|
end Pantograph.Test.Integration
|