Compare commits

...

15 Commits

14 changed files with 222 additions and 132 deletions

View File

@ -71,7 +71,7 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
Environment.addDecl args Environment.addDecl args
expr_echo (args: Protocol.ExprEcho): MainM (CR Protocol.ExprEchoResult) := do expr_echo (args: Protocol.ExprEcho): MainM (CR Protocol.ExprEchoResult) := do
let state ← get let state ← get
exprEcho args.expr state.options exprEcho args.expr args.type? state.options
options_set (args: Protocol.OptionsSet): MainM (CR Protocol.OptionsSetResult) := do options_set (args: Protocol.OptionsSet): MainM (CR Protocol.OptionsSetResult) := do
let state ← get let state ← get
let options := state.options let options := state.options
@ -93,11 +93,7 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
let state ← get let state ← get
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let expr?: Except _ GoalState ← runTermElabM (match args.expr, args.copyFrom with let expr?: Except _ GoalState ← runTermElabM (match args.expr, args.copyFrom with
| .some expr, .none => do | .some expr, .none => goalStartExpr expr
let expr ← match ← exprParse expr with
| .error e => return .error e
| .ok expr => pure $ expr
return .ok $ ← GoalState.create expr
| .none, .some copyFrom => | .none, .some copyFrom =>
(match env.find? <| copyFrom.toName with (match env.find? <| copyFrom.toName with
| .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}" | .none => return .error <| errorIndex s!"Symbol not found: {copyFrom}"
@ -152,7 +148,7 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
| .some branchId, .none => do | .some branchId, .none => do
match state.goalStates.find? branchId with match state.goalStates.find? branchId with
| .none => return .error $ errorIndex s!"Invalid state index {branchId}" | .none => return .error $ errorIndex s!"Invalid state index {branchId}"
| .some branch => pure $ target.continue branch | .some branch => pure $ goalContinue target branch
| .none, .some goals => | .none, .some goals =>
pure $ goalResume target goals pure $ goalResume target goals
| _, _ => return .error <| errorI "arguments" "Exactly one of {branch, goals} must be supplied" | _, _ => return .error <| errorI "arguments" "Exactly one of {branch, goals} must be supplied"

View File

@ -69,8 +69,8 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): CoreM (Pr
else pure (.none), else pure (.none),
module? := module? module? := module?
} }
let result := match info with let result match info with
| .inductInfo induct => { core with inductInfo? := .some { | .inductInfo induct => pure { core with inductInfo? := .some {
numParams := induct.numParams, numParams := induct.numParams,
numIndices := induct.numIndices, numIndices := induct.numIndices,
all := induct.all.toArray.map (·.toString), all := induct.all.toArray.map (·.toString),
@ -79,32 +79,38 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): CoreM (Pr
isReflexive := induct.isReflexive, isReflexive := induct.isReflexive,
isNested := induct.isNested, isNested := induct.isNested,
} } } }
| .ctorInfo ctor => { core with constructorInfo? := .some { | .ctorInfo ctor => pure { core with constructorInfo? := .some {
induct := ctor.induct.toString, induct := ctor.induct.toString,
cidx := ctor.cidx, cidx := ctor.cidx,
numParams := ctor.numParams, numParams := ctor.numParams,
numFields := ctor.numFields, numFields := ctor.numFields,
} } } }
| .recInfo r => { core with recursorInfo? := .some { | .recInfo r => pure { core with recursorInfo? := .some {
all := r.all.toArray.map (·.toString), all := r.all.toArray.map (·.toString),
numParams := r.numParams, numParams := r.numParams,
numIndices := r.numIndices, numIndices := r.numIndices,
numMotives := r.numMotives, numMotives := r.numMotives,
numMinors := r.numMinors, numMinors := r.numMinors,
rules := ← r.rules.toArray.mapM (λ rule => do
pure {
ctor := rule.ctor.toString,
nFields := rule.nfields,
rhs := ← (serialize_expression options rule.rhs).run',
})
k := r.k, k := r.k,
} } } }
| _ => core | _ => pure core
return .ok result return .ok result
def addDecl (args: Protocol.EnvAdd): CoreM (Protocol.CR Protocol.EnvAddResult) := do def addDecl (args: Protocol.EnvAdd): CoreM (Protocol.CR Protocol.EnvAddResult) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let tvM: Elab.TermElabM (Except String (Expr × Expr)) := do let tvM: Elab.TermElabM (Except String (Expr × Expr)) := do
let type ← match syntax_from_str env args.type with let type ← match parseTerm env args.type with
| .ok syn => do | .ok syn => do
match ← syntax_to_expr syn with match ← elabTerm syn with
| .error e => return .error e | .error e => return .error e
| .ok expr => pure expr | .ok expr => pure expr
| .error e => return .error e | .error e => return .error e
let value ← match syntax_from_str env args.value with let value ← match parseTerm env args.value with
| .ok syn => do | .ok syn => do
try try
let expr ← Elab.Term.elabTerm (stx := syn) (expectedType? := .some type) let expr ← Elab.Term.elabTerm (stx := syn) (expectedType? := .some type)

View File

@ -36,13 +36,15 @@ end Lean
namespace Pantograph namespace Pantograph
def defaultTermElabMContext: Lean.Elab.Term.Context := {
autoBoundImplicit := true,
declName? := some "_pantograph".toName,
errToSorry := false
}
def runMetaM { α } (metaM: Lean.MetaM α): Lean.CoreM α := def runMetaM { α } (metaM: Lean.MetaM α): Lean.CoreM α :=
metaM.run' metaM.run'
def runTermElabM { α } (termElabM: Lean.Elab.TermElabM α): Lean.CoreM α := def runTermElabM { α } (termElabM: Lean.Elab.TermElabM α): Lean.CoreM α :=
termElabM.run' (ctx := { termElabM.run' (ctx := defaultTermElabMContext) |>.run'
declName? := .none,
errToSorry := false,
}) |>.run'
def errorI (type desc: String): Protocol.InteractionError := { error := type, desc := desc } def errorI (type desc: String): Protocol.InteractionError := { error := type, desc := desc }
@ -109,21 +111,34 @@ def envAdd (name: String) (type: String) (value: String) (isTheorem: Bool):
Lean.CoreM (Protocol.CR Protocol.EnvAddResult) := Lean.CoreM (Protocol.CR Protocol.EnvAddResult) :=
Environment.addDecl { name, type, value, isTheorem } Environment.addDecl { name, type, value, isTheorem }
/-- This must be a TermElabM since the parsed expr contains extra information -/ def parseElabType (type: String): Lean.Elab.TermElabM (Protocol.CR Lean.Expr) := do
def exprParse (s: String): Lean.Elab.TermElabM (Protocol.CR Lean.Expr) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let syn ← match syntax_from_str env s with let syn ← match parseTerm env type with
| .error str => return .error $ errorI "parsing" str | .error str => return .error $ errorI "parsing" str
| .ok syn => pure syn | .ok syn => pure syn
match ← syntax_to_expr syn with match ← elabType syn with
| .error str => return .error $ errorI "elab" str | .error str => return .error $ errorI "elab" str
| .ok expr => return .ok expr | .ok expr => return .ok (← Lean.instantiateMVars expr)
/-- This must be a TermElabM since the parsed expr contains extra information -/
def parseElabExpr (expr: String) (expectedType?: Option String := .none): Lean.Elab.TermElabM (Protocol.CR Lean.Expr) := do
let env ← Lean.MonadEnv.getEnv
let expectedType? ← match ← expectedType?.mapM parseElabType with
| .none => pure $ .none
| .some (.ok t) => pure $ .some t
| .some (.error e) => return .error e
let syn ← match parseTerm env expr with
| .error str => return .error $ errorI "parsing" str
| .ok syn => pure syn
match ← elabTerm syn expectedType? with
| .error str => return .error $ errorI "elab" str
| .ok expr => return .ok (← Lean.instantiateMVars expr)
@[export pantograph_expr_echo_m] @[export pantograph_expr_echo_m]
def exprEcho (expr: String) (options: @&Protocol.Options): def exprEcho (expr: String) (expectedType?: Option String := .none) (options: @&Protocol.Options):
Lean.CoreM (Protocol.CR Protocol.ExprEchoResult) := do Lean.CoreM (Protocol.CR Protocol.ExprEchoResult) := do
let termElabM: Lean.Elab.TermElabM _ := do let termElabM: Lean.Elab.TermElabM _ := do
let expr ← match ← exprParse expr with let expr ← match ← parseElabExpr expr expectedType? with
| .error e => return .error e | .error e => return .error e
| .ok expr => pure expr | .ok expr => pure expr
try try
@ -139,7 +154,7 @@ def exprEcho (expr: String) (options: @&Protocol.Options):
@[export pantograph_goal_start_expr_m] @[export pantograph_goal_start_expr_m]
def goalStartExpr (expr: String): Lean.CoreM (Protocol.CR GoalState) := def goalStartExpr (expr: String): Lean.CoreM (Protocol.CR GoalState) :=
let termElabM: Lean.Elab.TermElabM _ := do let termElabM: Lean.Elab.TermElabM _ := do
let expr ← match ← exprParse expr with let expr ← match ← parseElabType expr with
| .error e => return .error e | .error e => return .error e
| .ok expr => pure $ expr | .ok expr => pure $ expr
return .ok $ ← GoalState.create expr return .ok $ ← GoalState.create expr

View File

@ -98,6 +98,7 @@ structure StatResult where
-- Return the type of an expression -- Return the type of an expression
structure ExprEcho where structure ExprEcho where
expr: String expr: String
type?: Option String
deriving Lean.FromJson deriving Lean.FromJson
structure ExprEchoResult where structure ExprEchoResult where
expr: Expression expr: Expression
@ -137,12 +138,20 @@ structure ConstructorInfo where
numParams: Nat numParams: Nat
numFields: Nat numFields: Nat
deriving Lean.ToJson deriving Lean.ToJson
/-- See `Lean/Declaration.lean` -/
structure RecursorRule where
ctor: String
nFields: Nat
rhs: Expression
deriving Lean.ToJson
structure RecursorInfo where structure RecursorInfo where
all: Array String all: Array String
numParams: Nat numParams: Nat
numIndices: Nat numIndices: Nat
numMotives: Nat numMotives: Nat
numMinors: Nat numMinors: Nat
rules: Array RecursorRule
k: Bool k: Bool
deriving Lean.ToJson deriving Lean.ToJson
structure EnvInspectResult where structure EnvInspectResult where

View File

@ -21,23 +21,22 @@ def unfoldAuxLemmas (e : Lean.Expr) : Lean.CoreM Lean.Expr := do
--- Input Functions --- --- Input Functions ---
/-- Read syntax object from string -/ /-- Read syntax object from string -/
def syntax_from_str (env: Environment) (s: String): Except String Syntax := def parseTerm (env: Environment) (s: String): Except String Syntax :=
Parser.runParserCategory Parser.runParserCategory
(env := env) (env := env)
(catName := `term) (catName := `term)
(input := s) (input := s)
(fileName := "<stdin>") (fileName := "<stdin>")
/-- Parse a syntax object. May generate additional metavariables! -/ /-- Parse a syntax object. May generate additional metavariables! -/
def syntax_to_expr_type (syn: Syntax): Elab.TermElabM (Except String Expr) := do def elabType (syn: Syntax): Elab.TermElabM (Except String Expr) := do
try try
let expr ← Elab.Term.elabType syn let expr ← Elab.Term.elabType syn
return .ok expr return .ok expr
catch ex => return .error (← ex.toMessageData.toString) catch ex => return .error (← ex.toMessageData.toString)
def syntax_to_expr (syn: Syntax): Elab.TermElabM (Except String Expr) := do def elabTerm (syn: Syntax) (expectedType? : Option Expr := .none): Elab.TermElabM (Except String Expr) := do
try try
let expr ← Elab.Term.elabTerm (stx := syn) (expectedType? := .none) let expr ← Elab.Term.elabTerm (stx := syn) expectedType?
return .ok expr return .ok expr
catch ex => return .error (← ex.toMessageData.toString) catch ex => return .error (← ex.toMessageData.toString)
@ -161,14 +160,12 @@ partial def serialize_expression_ast (expr: Expr) (sanitize: Bool := true): Meta
of_name (name: Name) := name_to_ast name sanitize of_name (name: Name) := name_to_ast name sanitize
def serialize_expression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.Expression := do def serialize_expression (options: @&Protocol.Options) (e: Expr): MetaM Protocol.Expression := do
let pp := toString (← Meta.ppExpr e) let pp?: Option String ← match options.printExprPretty with
let pp?: Option String := match options.printExprPretty with | true => pure $ .some $ toString $ ← Meta.ppExpr e
| true => .some pp | false => pure $ .none
| false => .none let sexp?: Option String ← match options.printExprAST with
let sexp: String ← serialize_expression_ast e | true => pure $ .some $ ← serialize_expression_ast e
let sexp?: Option String := match options.printExprAST with | false => pure $ .none
| true => .some sexp
| false => .none
return { return {
pp?, pp?,
sexp? sexp?

View File

@ -72,7 +72,7 @@ where the application of `assumption` should lead to a failure.
See `Pantograph/Protocol.lean` for a description of the parameters and return values in JSON. See `Pantograph/Protocol.lean` for a description of the parameters and return values in JSON.
- `reset`: Delete all cached expressions and proof trees - `reset`: Delete all cached expressions and proof trees
- `expr.echo {"expr": <expr>}`: Determine the type of an expression and round-trip it - `expr.echo {"expr": <expr>, "type": <optional expected type>}`: Determine the type of an expression and round-trip it
- `env.catalog`: Display a list of all safe Lean symbols in the current environment - `env.catalog`: Display a list of all safe Lean symbols in the current environment
- `env.inspect {"name": <name>, "value": <bool>}`: Show the type and package of a - `env.inspect {"name": <name>, "value": <bool>}`: Show the type and package of a
given symbol; If value flag is set, the value is printed or hidden. By default given symbol; If value flag is set, the value is printed or hidden. By default
@ -132,8 +132,5 @@ requires the presence of `lean-all`.
To run tests: To run tests:
``` sh ``` sh
nix build .#checks.${system}.test nix flake check
``` ```
For example, `${system}` could be `x86_64-linux`. Using `nix develop` drops the
current session into a development shell with fixed Lean version.

View File

@ -8,6 +8,7 @@ open Lean
namespace Pantograph namespace Pantograph
-- Auxiliary functions
namespace Protocol namespace Protocol
/-- Set internal names to "" -/ /-- Set internal names to "" -/
def Goal.devolatilize (goal: Goal): Goal := def Goal.devolatilize (goal: Goal): Goal :=
@ -24,6 +25,9 @@ def Goal.devolatilize (goal: Goal): Goal :=
deriving instance DecidableEq, Repr for Expression deriving instance DecidableEq, Repr for Expression
deriving instance DecidableEq, Repr for Variable deriving instance DecidableEq, Repr for Variable
deriving instance DecidableEq, Repr for Goal deriving instance DecidableEq, Repr for Goal
deriving instance DecidableEq, Repr for ExprEchoResult
deriving instance DecidableEq, Repr for InteractionError
deriving instance DecidableEq, Repr for Option
end Protocol end Protocol
def TacticResult.toString : TacticResult → String def TacticResult.toString : TacticResult → String
@ -34,12 +38,16 @@ def TacticResult.toString : TacticResult → String
| .parseError error => s!".parseError {error}" | .parseError error => s!".parseError {error}"
| .indexError index => s!".indexError {index}" | .indexError index => s!".indexError {index}"
namespace Test
def expectationFailure (desc: String) (error: String): LSpec.TestSeq := LSpec.test desc (LSpec.ExpectationFailure "ok _" error)
def assertUnreachable (message: String): LSpec.TestSeq := LSpec.check message false def assertUnreachable (message: String): LSpec.TestSeq := LSpec.check message false
open Lean def parseFailure (error: String) := expectationFailure "parse" error
def elabFailure (error: String) := expectationFailure "elab" error
def runCoreMSeq (env: Environment) (coreM: CoreM LSpec.TestSeq): IO LSpec.TestSeq := do def runCoreMSeq (env: Environment) (coreM: CoreM LSpec.TestSeq) (options: Array String := #[]): IO LSpec.TestSeq := do
let coreContext: Core.Context ← createCoreContext #[] let coreContext: Core.Context ← createCoreContext options
match ← (coreM.run' coreContext { env := env }).toBaseIO with match ← (coreM.run' coreContext { env := env }).toBaseIO with
| .error exception => | .error exception =>
return LSpec.test "Exception" (s!"internal exception #{← exception.toMessageData.toString}" = "") return LSpec.test "Exception" (s!"internal exception #{← exception.toMessageData.toString}" = "")
@ -47,14 +55,8 @@ def runCoreMSeq (env: Environment) (coreM: CoreM LSpec.TestSeq): IO LSpec.TestSe
def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSeq := def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSeq :=
runCoreMSeq env metaM.run' runCoreMSeq env metaM.run'
def runTermElabMInMeta { α } (termElabM: Lean.Elab.TermElabM α): Lean.MetaM α := def runTermElabMInMeta { α } (termElabM: Lean.Elab.TermElabM α): Lean.MetaM α :=
termElabM.run' (ctx := { termElabM.run' (ctx := Pantograph.defaultTermElabMContext)
declName? := .none,
errToSorry := false,
})
def defaultTermElabMContext: Lean.Elab.Term.Context := { end Test
declName? := some "_pantograph".toName,
errToSorry := false
}
end Pantograph end Pantograph

View File

@ -11,6 +11,7 @@ open Pantograph
deriving instance DecidableEq, Repr for Protocol.InductInfo deriving instance DecidableEq, Repr for Protocol.InductInfo
deriving instance DecidableEq, Repr for Protocol.ConstructorInfo deriving instance DecidableEq, Repr for Protocol.ConstructorInfo
deriving instance DecidableEq, Repr for Protocol.RecursorRule
deriving instance DecidableEq, Repr for Protocol.RecursorInfo deriving instance DecidableEq, Repr for Protocol.RecursorInfo
deriving instance DecidableEq, Repr for Protocol.EnvInspectResult deriving instance DecidableEq, Repr for Protocol.EnvInspectResult
@ -69,6 +70,7 @@ def test_inspect: IO LSpec.TestSeq := do
numIndices := 1, numIndices := 1,
numMotives := 1, numMotives := 1,
numMinors := 1, numMinors := 1,
rules := #[{ ctor := "Eq.refl", nFields := 0, rhs := { pp? := .some "fun {α} a motive refl => refl" } }]
k := true, k := true,
}), }),
("ForM.rec", ConstantCat.recursor { ("ForM.rec", ConstantCat.recursor {
@ -77,6 +79,7 @@ def test_inspect: IO LSpec.TestSeq := do
numIndices := 0, numIndices := 0,
numMotives := 1, numMotives := 1,
numMinors := 1, numMinors := 1,
rules := #[{ ctor := "ForM.mk", nFields := 1, rhs := { pp? := .some "fun m γ α motive mk forM => mk forM" } }]
k := false, k := false,
}) })
] ]
@ -94,10 +97,11 @@ def test_inspect: IO LSpec.TestSeq := do
) LSpec.TestSeq.done ) LSpec.TestSeq.done
runCoreMSeq env inner runCoreMSeq env inner
def suite: IO LSpec.TestSeq := do def suite: List (String × IO LSpec.TestSeq) :=
return LSpec.group "Environment" $ [
(LSpec.group "Catalog" (← test_catalog)) ++ ("Catalog", test_catalog),
(LSpec.group "Symbol visibility" (← test_symbol_visibility)) ++ ("Symbol Visibility", test_symbol_visibility),
(LSpec.group "Inspect" (← test_inspect)) ("Inspect", test_inspect),
]
end Pantograph.Test.Environment end Pantograph.Test.Environment

View File

@ -32,6 +32,16 @@ def subroutine_runner (steps: List (MainM LSpec.TestSeq)): IO LSpec.TestSeq := d
catch ex => catch ex =>
return LSpec.check s!"Uncaught IO exception: {ex.toString}" false return LSpec.check s!"Uncaught IO exception: {ex.toString}" false
def test_elab : IO LSpec.TestSeq :=
subroutine_runner [
subroutine_step "expr.echo"
[("expr", .str "λ {α : Sort (u + 1)} => List α")]
(Lean.toJson ({
type := { pp? := .some "{α : Type u} → Type u" },
expr := { pp? := .some "fun {α} => List α" }
}: Protocol.ExprEchoResult)),
]
def test_option_modify : IO LSpec.TestSeq := def test_option_modify : IO LSpec.TestSeq :=
let pp? := Option.some "∀ (n : Nat), n + 1 = n.succ" let pp? := Option.some "∀ (n : Nat), n + 1 = n.succ"
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)))" 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)))"
@ -73,22 +83,22 @@ def test_malformed_command : IO LSpec.TestSeq :=
] ]
def test_tactic : IO LSpec.TestSeq := def test_tactic : IO LSpec.TestSeq :=
let goal1: Protocol.Goal := { let goal1: Protocol.Goal := {
name := "_uniq.10", name := "_uniq.11",
target := { pp? := .some "∀ (q : Prop), x q → q x" }, target := { pp? := .some "∀ (q : Prop), x q → q x" },
vars := #[{ name := "_uniq.9", userName := "x", isInaccessible? := .some false, type? := .some { pp? := .some "Prop" }}], vars := #[{ name := "_uniq.10", userName := "x", isInaccessible? := .some false, type? := .some { pp? := .some "Prop" }}],
} }
let goal2: Protocol.Goal := { let goal2: Protocol.Goal := {
name := "_uniq.13", name := "_uniq.14",
target := { pp? := .some "x y → y x" }, target := { pp? := .some "x y → y x" },
vars := #[ vars := #[
{ name := "_uniq.9", userName := "x", isInaccessible? := .some false, type? := .some { pp? := .some "Prop" }}, { name := "_uniq.10", userName := "x", isInaccessible? := .some false, type? := .some { pp? := .some "Prop" }},
{ name := "_uniq.12", userName := "y", isInaccessible? := .some false, type? := .some { pp? := .some "Prop" }} { name := "_uniq.13", userName := "y", isInaccessible? := .some false, type? := .some { pp? := .some "Prop" }}
], ],
} }
subroutine_runner [ subroutine_runner [
subroutine_step "goal.start" subroutine_step "goal.start"
[("expr", .str "∀ (p q: Prop), p q → q p")] [("expr", .str "∀ (p q: Prop), p q → q p")]
(Lean.toJson ({stateId := 0, root := "_uniq.8"}: (Lean.toJson ({stateId := 0, root := "_uniq.9"}:
Protocol.GoalStartResult)), Protocol.GoalStartResult)),
subroutine_step "goal.tactic" subroutine_step "goal.tactic"
[("stateId", .num 0), ("goalId", .num 0), ("tactic", .str "intro x")] [("stateId", .num 0), ("goalId", .num 0), ("tactic", .str "intro x")]
@ -100,7 +110,7 @@ def test_tactic : IO LSpec.TestSeq :=
subroutine_step "goal.print" subroutine_step "goal.print"
[("stateId", .num 1)] [("stateId", .num 1)]
(Lean.toJson ({ (Lean.toJson ({
parent? := .some { pp? := .some "fun x => ?m.11 x" }, parent? := .some { pp? := .some "fun x => ?m.12 x" },
}: }:
Protocol.GoalPrintResult)), Protocol.GoalPrintResult)),
subroutine_step "goal.tactic" subroutine_step "goal.tactic"
@ -112,7 +122,7 @@ def test_tactic : IO LSpec.TestSeq :=
Protocol.GoalTacticResult)) Protocol.GoalTacticResult))
] ]
def test_env : IO LSpec.TestSeq := def test_env_add_inspect : IO LSpec.TestSeq :=
let name1 := "Pantograph.mystery" let name1 := "Pantograph.mystery"
let name2 := "Pantograph.mystery2" let name2 := "Pantograph.mystery2"
subroutine_runner [ subroutine_runner [
@ -148,13 +158,14 @@ def test_env : IO LSpec.TestSeq :=
Protocol.EnvInspectResult)) Protocol.EnvInspectResult))
] ]
def suite: IO LSpec.TestSeq := do def suite: List (String × IO LSpec.TestSeq) :=
[
return LSpec.group "Integration" $ ("Elab", test_elab),
(LSpec.group "Option modify" (← test_option_modify)) ++ ("Option modify", test_option_modify),
(LSpec.group "Malformed command" (← test_malformed_command)) ++ ("Malformed command", test_malformed_command),
(LSpec.group "Tactic" (← test_tactic)) ++ ("Tactic", test_tactic),
(LSpec.group "Env" (← test_env)) ("env.add env.inspect", test_env_add_inspect),
]
end Pantograph.Test.Integration end Pantograph.Test.Integration

38
Test/Library.lean Normal file
View File

@ -0,0 +1,38 @@
import LSpec
import Lean
import Pantograph.Library
import Test.Common
open Lean
open Pantograph
namespace Pantograph.Test.Library
def test_expr_echo (env: Environment): IO LSpec.TestSeq := do
let inner: CoreM LSpec.TestSeq := do
let prop_and_proof := "⟨∀ (x: Prop), x → x, λ (x: Prop) (h: x) => h⟩"
let tests := LSpec.TestSeq.done
let echoResult ← exprEcho prop_and_proof (options := {})
let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some {
type := { pp? := "?m.2" }, expr := { pp? := "?m.3" }
}))
let echoResult ← exprEcho prop_and_proof (expectedType? := .some "Σ' p:Prop, p") (options := { printExprAST := true })
let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some {
type := {
pp? := "(p : Prop) ×' p",
sexp? := "((:c PSigma) (:sort 0) (:lambda p (:sort 0) 0))",
},
expr := {
pp? := "⟨∀ (x : Prop), x → x, fun x h => h⟩",
sexp? := "((:c PSigma.mk) (:sort 0) (:lambda p (:sort 0) 0) (:forall x (:sort 0) (:forall _ 0 1)) (:lambda x (:sort 0) (:lambda h 0 0)))",
}
}))
return tests
runCoreMSeq env (options := #["pp.proofs.threshold=100"]) inner
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("expr_echo", test_expr_echo env),
]
end Pantograph.Test.Library

View File

@ -1,21 +1,53 @@
import LSpec import LSpec
import Test.Environment import Test.Environment
import Test.Metavar
import Test.Integration import Test.Integration
import Test.Library
import Test.Metavar
import Test.Proofs import Test.Proofs
import Test.Serial import Test.Serial
-- Test running infrastructure
namespace Pantograph.Test
def addPrefix (pref: String) (tests: List (String × α)): List (String × α) :=
tests.map (λ (name, x) => (pref ++ "/" ++ name, x))
/-- Runs test in parallel. Filters test name if given -/
def runTestGroup (filter: Option String) (tests: List (String × IO LSpec.TestSeq)): IO LSpec.TestSeq := do
let tests: List (String × IO LSpec.TestSeq) := match filter with
| .some filter => tests.filter (λ (name, _) => filter.isPrefixOf name)
| .none => tests
let tasks: List (String × Task _) ← tests.mapM (λ (name, task) => do
return (name, ← EIO.asTask task))
let all := tasks.foldl (λ acc (name, task) =>
let v: Except IO.Error LSpec.TestSeq := Task.get task
match v with
| .ok case => acc ++ (LSpec.group name case)
| .error e => acc ++ (expectationFailure name e.toString)
) LSpec.TestSeq.done
return all
end Pantograph.Test
open Pantograph.Test open Pantograph.Test
def main := do /-- Main entry of tests; Provide an argument to filter tests by prefix -/
def main (args: List String) := do
let name_filter := args.head?
Lean.initSearchPath (← Lean.findSysroot) Lean.initSearchPath (← Lean.findSysroot)
let env_default: Lean.Environment ← Lean.importModules
(imports := #[`Init])
(opts := {})
(trustLevel := 1)
let suites := [ let suites: List (String × List (String × IO LSpec.TestSeq)) := [
Metavar.suite, ("Environment", Environment.suite),
Integration.suite, ("Integration", Integration.suite),
Proofs.suite, ("Library", Library.suite env_default),
Serial.suite, ("Metavar", Metavar.suite env_default),
Environment.suite ("Proofs", Proofs.suite env_default),
("Serial", Serial.suite env_default),
] ]
let all ← suites.foldlM (λ acc m => do pure $ acc ++ (← m)) LSpec.TestSeq.done let tests: List (String × IO LSpec.TestSeq) := suites.foldl (λ acc (name, suite) => acc ++ (addPrefix name suite)) []
LSpec.lspecIO $ all LSpec.lspecIO (← runTestGroup name_filter tests)

View File

@ -17,12 +17,12 @@ def addTest (test: LSpec.TestSeq): TestM Unit := do
def test_instantiate_mvar: TestM Unit := do def test_instantiate_mvar: TestM Unit := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let value := "@Nat.le_trans 2 2 5 (@of_eq_true (@LE.le Nat instLENat 2 2) (@eq_true (@LE.le Nat instLENat 2 2) (@Nat.le_refl 2))) (@of_eq_true (@LE.le Nat instLENat 2 5) (@eq_true_of_decide (@LE.le Nat instLENat 2 5) (@Nat.decLe 2 5) (@Eq.refl Bool Bool.true)))" let value := "@Nat.le_trans 2 2 5 (@of_eq_true (@LE.le Nat instLENat 2 2) (@eq_true (@LE.le Nat instLENat 2 2) (@Nat.le_refl 2))) (@of_eq_true (@LE.le Nat instLENat 2 5) (@eq_true_of_decide (@LE.le Nat instLENat 2 5) (@Nat.decLe 2 5) (@Eq.refl Bool Bool.true)))"
let syn ← match syntax_from_str env value with let syn ← match parseTerm env value with
| .ok s => pure $ s | .ok s => pure $ s
| .error e => do | .error e => do
addTest $ assertUnreachable e addTest $ assertUnreachable e
return () return ()
let expr ← match ← syntax_to_expr syn with let expr ← match ← elabTerm syn with
| .ok expr => pure $ expr | .ok expr => pure $ expr
| .error e => do | .error e => do
addTest $ assertUnreachable e addTest $ assertUnreachable e
@ -36,14 +36,14 @@ def test_instantiate_mvar: TestM Unit := do
def startProof (expr: String): TestM (Option GoalState) := do def startProof (expr: String): TestM (Option GoalState) := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let syn? := syntax_from_str env expr let syn? := parseTerm env expr
addTest $ LSpec.check s!"Parsing {expr}" (syn?.isOk) addTest $ LSpec.check s!"Parsing {expr}" (syn?.isOk)
match syn? with match syn? with
| .error error => | .error error =>
IO.println error IO.println error
return Option.none return Option.none
| .ok syn => | .ok syn =>
let expr? ← syntax_to_expr_type syn let expr? ← elabType syn
addTest $ LSpec.check s!"Elaborating" expr?.isOk addTest $ LSpec.check s!"Elaborating" expr?.isOk
match expr? with match expr? with
| .error error => | .error error =>
@ -268,11 +268,7 @@ def test_partial_continuation: TestM Unit := do
return () return ()
def suite: IO LSpec.TestSeq := do def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
let env: Lean.Environment ← Lean.importModules
(imports := #["Init"].map (λ str => { module := str.toName, runtimeOnly := false }))
(opts := {})
(trustLevel := 1)
let tests := [ let tests := [
("Instantiate", test_instantiate_mvar), ("Instantiate", test_instantiate_mvar),
("2 < 5", test_m_couple), ("2 < 5", test_m_couple),
@ -280,11 +276,6 @@ def suite: IO LSpec.TestSeq := do
("Proposition Generation", test_proposition_generation), ("Proposition Generation", test_proposition_generation),
("Partial Continuation", test_partial_continuation) ("Partial Continuation", test_partial_continuation)
] ]
let tests ← tests.foldlM (fun acc tests => do tests.map (fun (name, test) => (name, proofRunner env test))
let (name, tests) := tests
let tests ← proofRunner env tests
return acc ++ (LSpec.group name tests)) LSpec.TestSeq.done
return LSpec.group "Metavar" tests
end Pantograph.Test.Metavar end Pantograph.Test.Metavar

View File

@ -32,14 +32,14 @@ def startProof (start: Start): TestM (Option GoalState) := do
| .none => | .none =>
return Option.none return Option.none
| .expr expr => | .expr expr =>
let syn? := syntax_from_str env expr let syn? := parseTerm env expr
addTest $ LSpec.check s!"Parsing {expr}" (syn?.isOk) addTest $ LSpec.check s!"Parsing {expr}" (syn?.isOk)
match syn? with match syn? with
| .error error => | .error error =>
IO.println error IO.println error
return Option.none return Option.none
| .ok syn => | .ok syn =>
let expr? ← syntax_to_expr_type syn let expr? ← elabType syn
addTest $ LSpec.check s!"Elaborating" expr?.isOk addTest $ LSpec.check s!"Elaborating" expr?.isOk
match expr? with match expr? with
| .error error => | .error error =>
@ -293,12 +293,7 @@ def proof_or_comm: TestM Unit := do
] ]
} }
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
def suite: IO LSpec.TestSeq := do
let env: Lean.Environment ← Lean.importModules
(imports := #[{ module := "Init".toName, runtimeOnly := false}])
(opts := {})
(trustLevel := 1)
let tests := [ let tests := [
("Nat.add_comm", proof_nat_add_comm false), ("Nat.add_comm", proof_nat_add_comm false),
("Nat.add_comm manual", proof_nat_add_comm true), ("Nat.add_comm manual", proof_nat_add_comm true),
@ -306,11 +301,6 @@ def suite: IO LSpec.TestSeq := do
("arithmetic", proof_arith), ("arithmetic", proof_arith),
("Or.comm", proof_or_comm) ("Or.comm", proof_or_comm)
] ]
let tests ← tests.foldlM (fun acc tests => do tests.map (fun (name, test) => (name, proofRunner env test))
let (name, tests) := tests
let tests ← proofRunner env tests
return acc ++ (LSpec.group name tests)) LSpec.TestSeq.done
return LSpec.group "Proofs" tests
end Pantograph.Test.Proofs end Pantograph.Test.Proofs

View File

@ -51,12 +51,18 @@ def test_sexp_of_symbol (env: Environment): IO LSpec.TestSeq := do
def test_sexp_of_expr (env: Environment): IO LSpec.TestSeq := do def test_sexp_of_expr (env: Environment): IO LSpec.TestSeq := do
let entries: List (String × String) := [ 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: 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))") ("λ x: Array Nat => x.data", "(:lambda x ((:c Array) (:c Nat)) ((:c Array.data) (:c Nat) 0))"),
-- This tests `autoBoundImplicit`
("λ {α : Sort (u + 1)} => List α", "(:lambda α (:sort (+ u 1)) ((:c List) 0) :implicit)"),
] ]
let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (source, target) => do let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (source, target) => do
let env ← MonadEnv.getEnv let env ← MonadEnv.getEnv
let s := syntax_from_str env source |>.toOption |>.get! let s ← match parseTerm env source with
let expr := (← syntax_to_expr s) |>.toOption |>.get! | .ok s => pure s
| .error e => return parseFailure e
let expr ← match (← elabTerm s) with
| .ok expr => pure expr
| .error e => return elabFailure e
let test := LSpec.check source ((← serialize_expression_ast expr) = target) let test := LSpec.check source ((← serialize_expression_ast expr) = target)
return LSpec.TestSeq.append suites test) LSpec.TestSeq.done return LSpec.TestSeq.append suites test) LSpec.TestSeq.done
let metaM := termElabM.run' (ctx := defaultTermElabMContext) let metaM := termElabM.run' (ctx := defaultTermElabMContext)
@ -67,22 +73,18 @@ def test_instance (env: Environment): IO LSpec.TestSeq := do
let metaM: MetaM LSpec.TestSeq := do let metaM: MetaM LSpec.TestSeq := do
let env ← MonadEnv.getEnv let env ← MonadEnv.getEnv
let source := "λ x y: Nat => HAdd.hAdd Nat Nat Nat (instHAdd Nat instAddNat) x y" let source := "λ x y: Nat => HAdd.hAdd Nat Nat Nat (instHAdd Nat instAddNat) x y"
let s := syntax_from_str env source |>.toOption |>.get! let s := parseTerm env source |>.toOption |>.get!
let _expr := (← runTermElabMInMeta <| syntax_to_expr s) |>.toOption |>.get! let _expr := (← runTermElabMInMeta <| elabTerm s) |>.toOption |>.get!
return LSpec.TestSeq.done return LSpec.TestSeq.done
runMetaMSeq env metaM runMetaMSeq env metaM
def suite: IO LSpec.TestSeq := do def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
let env: Environment ← importModules [
(imports := #["Init"].map (λ str => { module := str.toName, runtimeOnly := false })) ("name_to_ast", do pure test_name_to_ast),
(opts := {}) ("Expression binder", test_expr_to_binder env),
(trustLevel := 1) ("Sexp from symbol", test_sexp_of_symbol env),
("Sexp from expr", test_sexp_of_expr env),
return LSpec.group "Serialization" $ ("Instance", test_instance env),
(LSpec.group "name_to_ast" test_name_to_ast) ++ ]
(LSpec.group "Expression binder" (← test_expr_to_binder env)) ++
(LSpec.group "Sexp from symbol" (← test_sexp_of_symbol env)) ++
(LSpec.group "Sexp from expr" (← test_sexp_of_expr env)) ++
(LSpec.group "Instance" (← test_instance env))
end Pantograph.Test.Serial end Pantograph.Test.Serial