Pantograph/Test/Tactic/MotivatedApply.lean

105 lines
3.8 KiB
Plaintext
Raw Normal View History

2024-04-20 13:09:41 -07:00
import LSpec
import Lean
2024-04-22 00:11:41 -07:00
import Test.Common
2024-04-20 13:09:41 -07:00
open Lean
2024-04-22 00:11:41 -07:00
open Pantograph
2024-04-20 13:09:41 -07:00
namespace Pantograph.Test.Tactic.MotivatedApply
2024-04-22 00:11:41 -07:00
def valueAndType (recursor: String): MetaM (Expr × Expr) := do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := recursor)
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
runTermElabMInMeta do
let recursor ← Elab.Term.elabTerm (stx := recursor) .none
let recursorType ← Meta.inferType recursor
return (recursor, recursorType)
def test_type_extract (env: Environment): IO LSpec.TestSeq :=
runMetaMSeq env do
let mut tests := LSpec.TestSeq.done
let (recursor, recursorType) ← valueAndType "@Nat.brecOn"
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
2024-05-05 10:36:43 -07:00
def test_nat_brec_on (env: Environment): IO LSpec.TestSeq :=
2024-04-22 00:11:41 -07:00
let expr := "λ (n t: Nat) => n + 0 = n"
runMetaMSeq env do
let (expr, exprType) ← valueAndType 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",
2024-04-22 09:52:13 -07:00
"?motive ?m.69 = (n + 0 = n)",
2024-04-22 00:11:41 -07:00
])
tests := tests ++ test
return tests
2024-05-05 10:36:43 -07:00
def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
runMetaMSeq env do
let (expr, exprType) ← valueAndType 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.92",
"List ?m.94 → Prop",
"List ?m.94",
"∀ (t : List ?m.94), List.below t → ?motive t",
"?motive ?m.96 = (l ++ [] = [] ++ l)",
])
tests := tests ++ test
return tests
2024-04-20 13:09:41 -07:00
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
2024-04-22 00:11:41 -07:00
[
("type_extract", test_type_extract env),
2024-05-05 10:36:43 -07:00
("nat_brec_on", test_nat_brec_on env),
("list_brec_on", test_list_brec_on env),
2024-04-22 00:11:41 -07:00
]
2024-04-20 13:09:41 -07:00
end Pantograph.Test.Tactic.MotivatedApply