feat: Prograde tactics #83
|
@ -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
|
||||
|
|
|
@ -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),
|
||||
|
|
|
@ -7,16 +7,13 @@ 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)) =
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.30"),
|
||||
(`a₁, "?α"),
|
||||
|
@ -25,26 +22,20 @@ def test_congr_arg_list (env: Environment): IO LSpec.TestSeq :=
|
|||
(`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 :=
|
||||
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)) =
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.70"),
|
||||
(`a₁, "?α"),
|
||||
|
@ -53,18 +44,13 @@ def test_congr_arg (env: Environment): IO LSpec.TestSeq :=
|
|||
(`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 :=
|
||||
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)) =
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.159"),
|
||||
(`f₁, "?α → Nat"),
|
||||
|
@ -73,18 +59,13 @@ def test_congr_fun (env: Environment): IO LSpec.TestSeq :=
|
|||
(`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 :=
|
||||
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)) =
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.10"),
|
||||
(`f₁, "?α → Nat"),
|
||||
|
@ -95,15 +76,13 @@ def test_congr (env: Environment): IO LSpec.TestSeq :=
|
|||
(`h₂, "?a₁ = ?a₂"),
|
||||
(`conduit, "(?f₁ ?a₁ = ?f₂ ?a₂) = (a = b)"),
|
||||
])
|
||||
tests := tests ++ test
|
||||
return tests
|
||||
|
||||
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
|
||||
|
|
|
@ -7,25 +7,21 @@ 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
|
||||
def test_type_extract : TestT Elab.TermElabM Unit := do
|
||||
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" =
|
||||
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"
|
||||
tests := tests ++ LSpec.check "iMotive" (info.iMotive = 2)
|
||||
addTest $ LSpec.check "iMotive" (info.iMotive = 2)
|
||||
let motiveType := info.getMotiveType
|
||||
tests := tests ++ LSpec.check "motiveType" ("Nat → Sort ?u.1" =
|
||||
addTest $ LSpec.check "motiveType" ("Nat → Sort ?u.1" =
|
||||
(← exprToStr motiveType))
|
||||
return tests
|
||||
|
||||
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
|
||||
|
@ -35,25 +31,21 @@ def test_nat_brec_on (env: Environment): IO LSpec.TestSeq :=
|
|||
(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))) =
|
||||
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)",
|
||||
])
|
||||
tests := tests ++ test
|
||||
return tests
|
||||
addTest test
|
||||
|
||||
def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
|
||||
def test_list_brec_on : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
|
||||
runMetaMSeq env do
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
|
@ -63,13 +55,11 @@ def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
|
|||
(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))) =
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
|
||||
[
|
||||
"Type ?u.90",
|
||||
"List ?m.92 → Prop",
|
||||
|
@ -77,12 +67,9 @@ def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
|
|||
"∀ (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
|
||||
def test_partial_motive_instantiation : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n t: Nat) => n + 0 = n"
|
||||
runMetaMSeq env $ runTermElabMInMeta do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
|
@ -92,13 +79,12 @@ def test_partial_motive_instantiation (env: Environment): IO LSpec.TestSeq := do
|
|||
| .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))) =
|
||||
addTest $ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
|
||||
[
|
||||
"Nat → Prop",
|
||||
"Nat",
|
||||
|
@ -106,25 +92,22 @@ def test_partial_motive_instantiation (env: Environment): IO LSpec.TestSeq := do
|
|||
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}"))
|
||||
addTest $ (LSpec.check "goal name" (major.name.toString = s!"_uniq.{majorId}"))
|
||||
|
||||
-- Assign motive to `λ x => x + _`
|
||||
let motive_assign ← parseSentence "λ (x: Nat) => @Nat.add x + 0 = _"
|
||||
motive.assign motive_assign
|
||||
|
||||
let test ← conduit.withContext do
|
||||
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)")
|
||||
tests := tests ++ test
|
||||
|
||||
return tests
|
||||
|
||||
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
|
||||
|
|
|
@ -7,9 +7,8 @@ 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
|
||||
|
@ -19,20 +18,14 @@ def test_nat (env: Environment): IO LSpec.TestSeq :=
|
|||
(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
|
||||
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
|
||||
|
@ -42,21 +35,17 @@ def test_nat_fail (env: Environment): IO LSpec.TestSeq :=
|
|||
(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"
|
||||
let _ ← runTacticOnMVar tactic target.mvarId!
|
||||
addTest $ assertUnreachable "Tactic should fail"
|
||||
catch _ =>
|
||||
tests := tests ++ LSpec.check "Tactic should fail" true
|
||||
return tests
|
||||
return tests
|
||||
addTest $ LSpec.check "Tactic should fail" true
|
||||
|
||||
def test_list (env: Environment): IO LSpec.TestSeq :=
|
||||
def test_list : TestT Elab.TermElabM Unit := do
|
||||
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
|
||||
|
@ -66,22 +55,18 @@ def test_list (env: Environment): IO LSpec.TestSeq :=
|
|||
(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
|
||||
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
|
||||
|
|
|
@ -7,9 +7,8 @@ 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
|
||||
|
@ -19,31 +18,28 @@ def test_eval (env: Environment): IO LSpec.TestSeq :=
|
|||
(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 == {
|
||||
addTest $ LSpec.test "goals before" ((← toCondensedGoal goal.mvarId!).devolatilize == {
|
||||
context := #[
|
||||
{ userName := `p, type := .sort 0 },
|
||||
{ userName := `q, type := .sort 0 },
|
||||
{ userName := `h, type := h}
|
||||
cdeclOf `p (.sort 0),
|
||||
cdeclOf `q (.sort 0),
|
||||
cdeclOf `h 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 == {
|
||||
let [newGoal] ← runTacticOnMVar tactic goal.mvarId! | panic! "Incorrect goal number"
|
||||
addTest $ LSpec.test "goals after" ((← toCondensedGoal newGoal).devolatilize == {
|
||||
context := #[
|
||||
{ userName := `p, type := .sort 0 },
|
||||
{ userName := `q, type := .sort 0 },
|
||||
{ userName := `h, type := h},
|
||||
cdeclOf `p (.sort 0),
|
||||
cdeclOf `q (.sort 0),
|
||||
cdeclOf `h h,
|
||||
{
|
||||
userName := `h2,
|
||||
type := mkOr h m,
|
||||
|
@ -52,12 +48,72 @@ def test_eval (env: Environment): IO LSpec.TestSeq :=
|
|||
],
|
||||
target,
|
||||
})
|
||||
tests := tests ++ test
|
||||
return tests
|
||||
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
|
||||
|
|
Loading…
Reference in New Issue