refactor: Simplify proof test infrastructure

tactic/eval
Leni Aniva 2024-06-27 14:34:21 -04:00
parent 2d2ff24017
commit fc0d872343
Signed by: aniva
GPG Key ID: 4D9B1C8D10EA4C50
6 changed files with 349 additions and 378 deletions

View File

@ -89,10 +89,12 @@ def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSe
runCoreMSeq env metaM.run'
def runTermElabMInMeta { α } (termElabM: Lean.Elab.TermElabM α): Lean.MetaM α :=
termElabM.run' (ctx := Pantograph.defaultTermElabMContext)
def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq): IO LSpec.TestSeq :=
runMetaMSeq env $ termElabM.run' (ctx := Pantograph.defaultTermElabMContext)
def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e
def parseSentence (s: String): MetaM Expr := do
def parseSentence (s: String): Elab.TermElabM Expr := do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
@ -100,7 +102,7 @@ def parseSentence (s: String): MetaM Expr := do
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
runTermElabMInMeta $ Elab.Term.elabTerm (stx := recursor) .none
Elab.Term.elabTerm (stx := recursor) .none
def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
@ -110,6 +112,36 @@ def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do
let t ← exprToStr (← mvarId.getType)
return (name, t)
-- Monadic testing
abbrev TestT := StateT LSpec.TestSeq
def addTest [Monad m] (test: LSpec.TestSeq): TestT m Unit := do
set $ (← get) ++ test
def runTest [Monad m] (t: TestT m Unit): m LSpec.TestSeq :=
Prod.snd <$> t.run LSpec.TestSeq.done
def runTestTermElabM (env: Environment) (t: TestT Elab.TermElabM Unit):
IO LSpec.TestSeq :=
runTermElabMSeq env $ runTest t
def cdeclOf (userName: Name) (type: Expr): Condensed.LocalDecl :=
{ userName, type }
def buildGoal (nameType: List (String × String)) (target: String) (userName?: Option String := .none):
Protocol.Goal :=
{
userName?,
target := { pp? := .some target},
vars := (nameType.map fun x => ({
userName := x.fst,
type? := .some { pp? := .some x.snd },
isInaccessible? := .some false
})).toArray
}
end Test
end Pantograph

View File

@ -357,69 +357,6 @@ def test_or_comm: TestM Unit := do
]
}
def test_have: TestM Unit := do
let state? ← startProof (.expr "∀ (p q: Prop), p → ((p q) (p q))")
let state0 ← match state? with
| .some state => pure state
| .none => do
addTest $ assertUnreachable "Goal could not parse"
return ()
let tactic := "intro p q h"
let state1 ← match ← state0.tryTactic (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
#[buildGoal [("p", "Prop"), ("q", "Prop"), ("h", "p")] "(p q) p q"])
let expr := "Or.inl (Or.inl h)"
let state2 ← match ← state1.tryAssign (goalId := 0) (expr := expr) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!":= {expr}" ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
#[])
let haveBind := "y"
let haveType := "p q"
let state2 ← match ← state1.tryHave (goalId := 0) (binderName := haveBind) (type := haveType) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!"have {haveBind}: {haveType}" ((← state2.serializeGoals (options := ← read)).map (·.devolatilize) =
#[
buildGoal [("p", "Prop"), ("q", "Prop"), ("h", "p")] "p q",
buildGoal [("p", "Prop"), ("q", "Prop"), ("h", "p"), ("y", "p q")] "(p q) p q"
])
let expr := "Or.inl h"
let state3 ← match ← state2.tryAssign (goalId := 0) (expr := expr) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!":= {expr}" ((← state3.serializeGoals (options := ← read)).map (·.devolatilize) =
#[])
let state2b ← match state3.continue state2 with
| .ok state => pure state
| .error e => do
addTest $ assertUnreachable e
return ()
let expr := "Or.inl y"
let state4 ← match ← state2b.tryAssign (goalId := 0) (expr := expr) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!":= {expr}" ((← state4.serializeGoals (options := ← read)).map (·.devolatilize) =
#[])
addTest $ LSpec.check "(4 root)" state4.rootExpr?.isSome
example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by
intro a b c1 c2 h
conv =>
@ -856,7 +793,6 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
("Nat.add_comm delta", test_delta_variable),
("arithmetic", test_arith),
("Or.comm", test_or_comm),
("have", test_have),
("conv", test_conv),
("calc", test_calc),
("let via assign", test_let false),

View File

@ -7,103 +7,82 @@ open Pantograph
namespace Pantograph.Test.Tactic.Congruence
def test_congr_arg_list (env: Environment): IO LSpec.TestSeq :=
def test_congr_arg_list : TestT Elab.TermElabM Unit := do
let expr := "λ {α} (l1 l2 : List α) (h: l1 = l2) => l1.reverse = l2.reverse"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let mut tests := LSpec.TestSeq.done
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let (newGoals, test) ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar Tactic.congruenceArg target.mvarId!
let test := LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.30"),
(`a₁, "?α"),
(`a₂, "?α"),
(`f, "?α → List α"),
(`h, "?a₁ = ?a₂"),
(`conduit, "(?f ?a₁ = ?f ?a₂) = (l1.reverse = l2.reverse)"),
])
return (newGoals, test)
tests := tests ++ test
let f := newGoals.get! 3
let h := newGoals.get! 4
let c := newGoals.get! 5
let results ← f.apply (← parseSentence "List.reverse")
tests := tests ++ (LSpec.check "apply" (results.length = 0))
tests := tests ++ (LSpec.check "h" ((← exprToStr $ ← h.getType) = "?a₁ = ?a₂"))
tests := tests ++ (LSpec.check "conduit" ((← exprToStr $ ← c.getType) = "(?a₁.reverse = ?a₂.reverse) = (l1.reverse = l2.reverse)"))
return tests
def test_congr_arg (env: Environment): IO LSpec.TestSeq :=
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.congruenceArg target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.30"),
(`a₁, "?α"),
(`a₂, "?α"),
(`f, "?α → List α"),
(`h, "?a₁ = ?a₂"),
(`conduit, "(?f ?a₁ = ?f ?a₂) = (l1.reverse = l2.reverse)"),
])
let f := newGoals.get! 3
let h := newGoals.get! 4
let c := newGoals.get! 5
let results ← f.apply (← parseSentence "List.reverse")
addTest $ LSpec.check "apply" (results.length = 0)
addTest $ LSpec.check "h" ((← exprToStr $ ← h.getType) = "?a₁ = ?a₂")
addTest $ LSpec.check "conduit" ((← exprToStr $ ← c.getType) = "(?a₁.reverse = ?a₂.reverse) = (l1.reverse = l2.reverse)")
def test_congr_arg : TestT Elab.TermElabM Unit := do
let expr := "λ (n m: Nat) (h: n = m) => n * n = m * m"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let mut tests := LSpec.TestSeq.done
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar Tactic.congruenceArg target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.70"),
(`a₁, "?α"),
(`a₂, "?α"),
(`f, "?α → Nat"),
(`h, "?a₁ = ?a₂"),
(`conduit, "(?f ?a₁ = ?f ?a₂) = (n * n = m * m)"),
])
tests := tests ++ test
return tests
def test_congr_fun (env: Environment): IO LSpec.TestSeq :=
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.congruenceArg target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.70"),
(`a₁, "?α"),
(`a₂, "?α"),
(`f, "?α → Nat"),
(`h, "?a₁ = ?a₂"),
(`conduit, "(?f ?a₁ = ?f ?a₂) = (n * n = m * m)"),
])
def test_congr_fun : TestT Elab.TermElabM Unit := do
let expr := "λ (n m: Nat) => (n + m) + (n + m) = (n + m) * 2"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let mut tests := LSpec.TestSeq.done
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar Tactic.congruenceFun target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.159"),
(`f₁, "?α → Nat"),
(`f₂, "?α → Nat"),
(`h, "?f₁ = ?f₂"),
(`a, "?α"),
(`conduit, "(?f₁ ?a = ?f₂ ?a) = (n + m + (n + m) = (n + m) * 2)"),
])
tests := tests ++ test
return tests
def test_congr (env: Environment): IO LSpec.TestSeq :=
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.congruenceFun target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.159"),
(`f₁, "?α → Nat"),
(`f₂, "?α → Nat"),
(`h, "?f₁ = ?f₂"),
(`a, "?α"),
(`conduit, "(?f₁ ?a = ?f₂ ?a) = (n + m + (n + m) = (n + m) * 2)"),
])
def test_congr : TestT Elab.TermElabM Unit := do
let expr := "λ (a b: Nat) => a = b"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let mut tests := LSpec.TestSeq.done
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar Tactic.congruence target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.10"),
(`f₁, "?α → Nat"),
(`f₂, "?α → Nat"),
(`a₁, "?α"),
(`a₂, "?α"),
(`h₁, "?f₁ = ?f₂"),
(`h₂, "?a₁ = ?a₂"),
(`conduit, "(?f₁ ?a₁ = ?f₂ ?a₂) = (a = b)"),
])
tests := tests ++ test
return tests
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let newGoals ← runTacticOnMVar Tactic.congruence target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
[
(`α, "Sort ?u.10"),
(`f₁, "?α → Nat"),
(`f₂, "?α → Nat"),
(`a₁, "?α"),
(`a₂, "?α"),
(`h₁, "?f₁ = ?f₂"),
(`h₂, "?a₁ = ?a₂"),
(`conduit, "(?f₁ ?a₁ = ?f₂ ?a₂) = (a = b)"),
])
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("congrArg List.reverse", test_congr_arg_list env),
("congrArg", test_congr_arg env),
("congrFun", test_congr_fun env),
("congr", test_congr env),
]
("congrArg List.reverse", test_congr_arg_list),
("congrArg", test_congr_arg),
("congrFun", test_congr_fun),
("congr", test_congr),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.Congruence

View File

@ -7,82 +7,23 @@ open Pantograph
namespace Pantograph.Test.Tactic.MotivatedApply
def test_type_extract (env: Environment): IO LSpec.TestSeq :=
runMetaMSeq env do
let mut tests := LSpec.TestSeq.done
let recursor ← parseSentence "@Nat.brecOn"
let recursorType ← Meta.inferType recursor
tests := tests ++ LSpec.check "recursorType" ("{motive : Nat → Sort ?u.1} → (t : Nat) → ((t : Nat) → Nat.below t → motive t) → motive t" =
(← exprToStr recursorType))
let info ← match Tactic.getRecursorInformation recursorType with
| .some info => pure info
| .none => throwError "Failed to extract recursor info"
tests := tests ++ LSpec.check "iMotive" (info.iMotive = 2)
let motiveType := info.getMotiveType
tests := tests ++ LSpec.check "motiveType" ("Nat → Sort ?u.1" =
(← exprToStr motiveType))
return tests
def test_type_extract : TestT Elab.TermElabM Unit := do
let recursor ← parseSentence "@Nat.brecOn"
let recursorType ← Meta.inferType recursor
addTest $ LSpec.check "recursorType" ("{motive : Nat → Sort ?u.1} → (t : Nat) → ((t : Nat) → Nat.below t → motive t) → motive t" =
(← exprToStr recursorType))
let info ← match Tactic.getRecursorInformation recursorType with
| .some info => pure info
| .none => throwError "Failed to extract recursor info"
addTest $ LSpec.check "iMotive" (info.iMotive = 2)
let motiveType := info.getMotiveType
addTest $ LSpec.check "motiveType" ("Nat → Sort ?u.1" =
(← exprToStr motiveType))
def test_nat_brec_on (env: Environment): IO LSpec.TestSeq :=
def test_nat_brec_on : TestT Elab.TermElabM Unit := do
let expr := "λ (n t: Nat) => n + 0 = n"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@Nat.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar tactic target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Nat → Prop",
"Nat",
"∀ (t : Nat), Nat.below t → ?motive t",
"?motive ?m.67 = (n + 0 = n)",
])
tests := tests ++ test
return tests
def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@List.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar tactic target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Type ?u.90",
"List ?m.92 → Prop",
"List ?m.92",
"∀ (t : List ?m.92), List.below t → ?motive t",
"?motive ?m.94 = (l ++ [] = [] ++ l)",
])
tests := tests ++ test
return tests
def test_partial_motive_instantiation (env: Environment): IO LSpec.TestSeq := do
let expr := "λ (n t: Nat) => n + 0 = n"
runMetaMSeq env $ runTermElabMInMeta do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
@ -90,41 +31,83 @@ def test_partial_motive_instantiation (env: Environment): IO LSpec.TestSeq := do
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
let majorId := 67
tests := tests ++ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Nat → Prop",
"Nat",
"∀ (t : Nat), Nat.below t → ?motive t",
s!"?motive ?m.{majorId} = (n + 0 = n)",
]))
let [motive, major, step, conduit] := newGoals | panic! "Incorrect goal number"
tests := tests ++ (LSpec.check "goal name" (major.name.toString = s!"_uniq.{majorId}"))
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
let test := LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Nat → Prop",
"Nat",
"∀ (t : Nat), Nat.below t → ?motive t",
"?motive ?m.67 = (n + 0 = n)",
])
addTest test
-- Assign motive to `λ x => x + _`
let motive_assign ← parseSentence "λ (x: Nat) => @Nat.add x + 0 = _"
motive.assign motive_assign
def test_list_brec_on : TestT Elab.TermElabM Unit := do
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@List.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Type ?u.90",
"List ?m.92 → Prop",
"List ?m.92",
"∀ (t : List ?m.92), List.below t → ?motive t",
"?motive ?m.94 = (l ++ [] = [] ++ l)",
])
let test ← conduit.withContext do
let t := toString (← Meta.ppExpr $ ← conduit.getType)
return LSpec.check "conduit" (t = s!"(?m.{majorId}.add + 0 = ?m.138 ?m.{majorId}) = (n + 0 = n)")
tests := tests ++ test
def test_partial_motive_instantiation : TestT Elab.TermElabM Unit := do
let expr := "λ (n t: Nat) => n + 0 = n"
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "@Nat.brecOn")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
let majorId := 67
addTest $ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[
"Nat → Prop",
"Nat",
"∀ (t : Nat), Nat.below t → ?motive t",
s!"?motive ?m.{majorId} = (n + 0 = n)",
]))
let [motive, major, step, conduit] := newGoals | panic! "Incorrect goal number"
addTest $ (LSpec.check "goal name" (major.name.toString = s!"_uniq.{majorId}"))
return tests
-- Assign motive to `λ x => x + _`
let motive_assign ← parseSentence "λ (x: Nat) => @Nat.add x + 0 = _"
motive.assign motive_assign
addTest $ ← conduit.withContext do
let t := toString (← Meta.ppExpr $ ← conduit.getType)
return LSpec.check "conduit" (t = s!"(?m.{majorId}.add + 0 = ?m.138 ?m.{majorId}) = (n + 0 = n)")
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("type_extract", test_type_extract env),
("Nat.brecOn", test_nat_brec_on env),
("List.brecOn", test_list_brec_on env),
("Nat.brecOn partial motive instantiation", test_partial_motive_instantiation env),
]
("type_extract", test_type_extract),
("Nat.brecOn", test_nat_brec_on),
("List.brecOn", test_list_brec_on),
("Nat.brecOn partial motive instantiation", test_partial_motive_instantiation),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.MotivatedApply

View File

@ -7,81 +7,66 @@ open Pantograph
namespace Pantograph.Test.Tactic.NoConfuse
def test_nat (env: Environment): IO LSpec.TestSeq :=
def test_nat : TestT Elab.TermElabM Unit := do
let expr := "λ (n: Nat) (h: 0 = n + 1) => False"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.noConfuse recursor
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar tactic target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[])
tests := tests ++ test
return tests
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.noConfuse recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = [])
def test_nat_fail (env: Environment): IO LSpec.TestSeq :=
def test_nat_fail : TestT Elab.TermElabM Unit := do
let expr := "λ (n: Nat) (h: n = n) => False"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
try
let tactic := Tactic.noConfuse recursor
let _ ← runTermElabMInMeta $ runTacticOnMVar tactic target.mvarId!
tests := tests ++ assertUnreachable "Tactic should fail"
catch _ =>
tests := tests ++ LSpec.check "Tactic should fail" true
return tests
return tests
def test_list (env: Environment): IO LSpec.TestSeq :=
let expr := "λ (l: List Nat) (h: [] = 1 :: l) => False"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
try
let tactic := Tactic.noConfuse recursor
let test ← runTermElabMInMeta do
let newGoals ← runTacticOnMVar tactic target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[])
tests := tests ++ test
return tests
let _ ← runTacticOnMVar tactic target.mvarId!
addTest $ assertUnreachable "Tactic should fail"
catch _ =>
addTest $ LSpec.check "Tactic should fail" true
def test_list : TestT Elab.TermElabM Unit := do
let expr := "λ (l: List Nat) (h: [] = 1 :: l) => False"
let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.noConfuse recursor
let newGoals ← runTacticOnMVar tactic target.mvarId!
addTest $ LSpec.check "goals"
((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = [])
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("Nat", test_nat env),
("Nat fail", test_nat_fail env),
("List", test_list env),
]
("Nat", test_nat),
("Nat fail", test_nat_fail),
("List", test_list),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.NoConfuse

View File

@ -7,57 +7,113 @@ open Pantograph
namespace Pantograph.Test.Tactic.Prograde
def test_eval (env: Environment): IO LSpec.TestSeq :=
def test_eval : TestT Elab.TermElabM Unit := do
let expr := "forall (p q : Prop) (h: p), And (Or p q) (Or p q)"
runMetaMSeq env do
let expr ← parseSentence expr
Meta.forallTelescope expr $ λ _ body => do
let e ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "Or.inl h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
let mut tests := LSpec.TestSeq.done
-- Apply the tactic
let goal ← Meta.mkFreshExprSyntheticOpaqueMVar body
let target: Expr := mkAnd
(mkOr (.fvar ⟨uniq 8⟩) (.fvar ⟨uniq 9⟩))
(mkOr (.fvar ⟨uniq 8⟩) (.fvar ⟨uniq 9⟩))
let h := .fvar ⟨uniq 8⟩
let test := LSpec.test "goals before" ((← toCondensedGoal goal.mvarId!).devolatilize == {
context := #[
{ userName := `p, type := .sort 0 },
{ userName := `q, type := .sort 0 },
{ userName := `h, type := h}
],
target,
})
tests := tests ++ test
let tactic := Tactic.evaluate `h2 e
let m := .mvar ⟨uniq 13⟩
let test ← runTermElabMInMeta do
let [goal] ← runTacticOnMVar tactic goal.mvarId! | panic! "Incorrect goal number"
pure $ LSpec.test "goals after" ((← toCondensedGoal goal).devolatilize == {
context := #[
{ userName := `p, type := .sort 0 },
{ userName := `q, type := .sort 0 },
{ userName := `h, type := h},
{
userName := `h2,
type := mkOr h m,
value? := .some $ mkApp3 (mkConst `Or.inl) h m (.fvar ⟨uniq 10⟩)
}
],
target,
})
tests := tests ++ test
return tests
let expr ← parseSentence expr
Meta.forallTelescope expr $ λ _ body => do
let e ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := "Or.inl h")
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
-- Apply the tactic
let goal ← Meta.mkFreshExprSyntheticOpaqueMVar body
let target: Expr := mkAnd
(mkOr (.fvar ⟨uniq 8⟩) (.fvar ⟨uniq 9⟩))
(mkOr (.fvar ⟨uniq 8⟩) (.fvar ⟨uniq 9⟩))
let h := .fvar ⟨uniq 8⟩
addTest $ LSpec.test "goals before" ((← toCondensedGoal goal.mvarId!).devolatilize == {
context := #[
cdeclOf `p (.sort 0),
cdeclOf `q (.sort 0),
cdeclOf `h h
],
target,
})
let tactic := Tactic.evaluate `h2 e
let m := .mvar ⟨uniq 13⟩
let [newGoal] ← runTacticOnMVar tactic goal.mvarId! | panic! "Incorrect goal number"
addTest $ LSpec.test "goals after" ((← toCondensedGoal newGoal).devolatilize == {
context := #[
cdeclOf `p (.sort 0),
cdeclOf `q (.sort 0),
cdeclOf `h h,
{
userName := `h2,
type := mkOr h m,
value? := .some $ mkApp3 (mkConst `Or.inl) h m (.fvar ⟨uniq 10⟩)
}
],
target,
})
addTest $ LSpec.test "assign" ((← getExprMVarAssignment? goal.mvarId!) == .some (.mvar newGoal))
def test_have : TestT Elab.TermElabM Unit := do
let rootExpr ← parseSentence "∀ (p q: Prop), p → ((p q) (p q))"
let state0 ← GoalState.create rootExpr
let tactic := "intro p q h"
let state1 ← match ← state0.tryTactic (goalId := 0) (tactic := tactic) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) =
#[buildGoal [("p", "Prop"), ("q", "Prop"), ("h", "p")] "(p q) p q"])
let expr := "Or.inl (Or.inl h)"
let state2 ← match ← state1.tryAssign (goalId := 0) (expr := expr) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!":= {expr}" ((← state2.serializeGoals).map (·.devolatilize) =
#[])
let haveBind := "y"
let haveType := "p q"
let state2 ← match ← state1.tryHave (goalId := 0) (binderName := haveBind) (type := haveType) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!"have {haveBind}: {haveType}" ((← state2.serializeGoals).map (·.devolatilize) =
#[
buildGoal [("p", "Prop"), ("q", "Prop"), ("h", "p")] "p q",
buildGoal [("p", "Prop"), ("q", "Prop"), ("h", "p"), ("y", "p q")] "(p q) p q"
])
let expr := "Or.inl h"
let state3 ← match ← state2.tryAssign (goalId := 0) (expr := expr) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!":= {expr}" ((← state3.serializeGoals).map (·.devolatilize) =
#[])
let state2b ← match state3.continue state2 with
| .ok state => pure state
| .error e => do
addTest $ assertUnreachable e
return ()
let expr := "Or.inl y"
let state4 ← match ← state2b.tryAssign (goalId := 0) (expr := expr) with
| .success state => pure state
| other => do
addTest $ assertUnreachable $ other.toString
return ()
addTest $ LSpec.check s!":= {expr}" ((← state4.serializeGoals).map (·.devolatilize) =
#[])
addTest $ LSpec.check "(4 root)" state4.rootExpr?.isSome
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("eval", test_eval env),
]
("eval", test_eval),
("have", test_have),
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
end Pantograph.Test.Tactic.Prograde