Compare commits

..

2 Commits

Author SHA1 Message Date
Leni Aniva 75df7268c5
test: Simplify testing structure for tactics 2024-05-20 11:55:38 -07:00
Leni Aniva 92acf7782c
test: CongruenceArg tactic 2024-05-20 11:51:35 -07:00
7 changed files with 119 additions and 84 deletions

View File

@ -6,73 +6,80 @@ namespace Pantograph.Tactic
def congruenceArg: Elab.Tactic.TacticM Unit := do def congruenceArg: Elab.Tactic.TacticM Unit := do
let goal ← Elab.Tactic.getMainGoal let goal ← Elab.Tactic.getMainGoal
let .some (beta, _, _) := (← goal.getType).eq? | throwError "Goal is not an Eq" let .some (β, _, _) := (← goal.getType).eq? | throwError "Goal is not an Eq"
-- Create the descendant goals let userName := (← goal.getDecl).userName
let nextGoals ← goal.withContext do let nextGoals ← goal.withContext do
let u ← Meta.mkFreshLevelMVar let u ← Meta.mkFreshLevelMVar
let alpha ← Meta.mkFreshExprMVar (.some $ mkSort u) .natural .anonymous let α ← Meta.mkFreshExprMVar (.some $ mkSort u)
let f ← Meta.mkFreshExprMVar (.some <| .forallE .anonymous alpha beta .default) .natural (userName := userName ++ `α)
.synthetic (userName := goal.name ++ `f) let f ← Meta.mkFreshExprMVar (.some <| .forallE .anonymous α β .default)
let a₁ ← Meta.mkFreshExprMVar (.some alpha) .synthetic (userName := userName ++ `f)
.synthetic (userName := goal.name ++ `a₁) let a₁ ← Meta.mkFreshExprMVar (.some α)
let a₂ ← Meta.mkFreshExprMVar (.some alpha) .synthetic (userName := userName ++ `a₁)
.synthetic (userName := goal.name ++ `a₂) let a₂ ← Meta.mkFreshExprMVar (.some α)
let h ← Meta.mkEq a₁ a₂ .synthetic (userName := userName ++ `a₂)
let h ← Meta.mkFreshExprMVar (.some $ ← Meta.mkEq a₁ a₂)
.synthetic (userName := userName ++ `h)
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f a₁) (.app f a₂)) (← goal.getType) let conduitType ← Meta.mkEq (← Meta.mkEq (.app f a₁) (.app f a₂)) (← goal.getType)
let conduit ← Meta.mkFreshExprMVar conduitType let conduit ← Meta.mkFreshExprMVar conduitType
.synthetic (userName := goal.name ++ `conduit) .synthetic (userName := userName ++ `conduit)
goal.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrArg f h) goal.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrArg f h)
return [alpha, a₁, a₂, f, h, conduit] return [α, a₁, a₂, f, h, conduit]
Elab.Tactic.setGoals <| nextGoals.map (·.mvarId!) Elab.Tactic.setGoals <| nextGoals.map (·.mvarId!)
def congruenceFun: Elab.Tactic.TacticM Unit := do def congruenceFun: Elab.Tactic.TacticM Unit := do
let goal ← Elab.Tactic.getMainGoal let goal ← Elab.Tactic.getMainGoal
let .some (beta, _, _) := (← goal.getType).eq? | throwError "Goal is not an Eq" let .some (β, _, _) := (← goal.getType).eq? | throwError "Goal is not an Eq"
-- Create the descendant goals let userName := (← goal.getDecl).userName
let nextGoals ← goal.withContext do let nextGoals ← goal.withContext do
let u ← Meta.mkFreshLevelMVar let u ← Meta.mkFreshLevelMVar
let alpha ← Meta.mkFreshExprMVar (.some $ mkSort u) .natural .anonymous let α ← Meta.mkFreshExprMVar (.some $ mkSort u)
let fType := .forallE .anonymous alpha beta .default .natural (userName := userName ++ `α)
let fType := .forallE .anonymous α β .default
let f₁ ← Meta.mkFreshExprMVar (.some fType) let f₁ ← Meta.mkFreshExprMVar (.some fType)
.synthetic (userName := goal.name ++ `f₁) .synthetic (userName := userName ++ `f₁)
let f₂ ← Meta.mkFreshExprMVar (.some fType) let f₂ ← Meta.mkFreshExprMVar (.some fType)
.synthetic (userName := goal.name ++ `f₂) .synthetic (userName := userName ++ `f₂)
let a ← Meta.mkFreshExprMVar (.some alpha) let a ← Meta.mkFreshExprMVar (.some α)
.synthetic (userName := goal.name ++ `a) .synthetic (userName := userName ++ `a)
let h ← Meta.mkEq f₁ f₂ let h ← Meta.mkFreshExprMVar (.some $ ← Meta.mkEq f₁ f₂)
.synthetic (userName := userName ++ `h)
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a) (.app f₂ a)) (← goal.getType) let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a) (.app f₂ a)) (← goal.getType)
let conduit ← Meta.mkFreshExprMVar conduitType let conduit ← Meta.mkFreshExprMVar conduitType
.synthetic (userName := goal.name ++ `conduit) .synthetic (userName := userName ++ `conduit)
goal.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrFun h a) goal.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrFun h a)
return [alpha, f₁, f₂, h, a, conduit] return [α, f₁, f₂, h, a, conduit]
Elab.Tactic.setGoals <| nextGoals.map (·.mvarId!) Elab.Tactic.setGoals <| nextGoals.map (·.mvarId!)
def congruence: Elab.Tactic.TacticM Unit := do def congruence: Elab.Tactic.TacticM Unit := do
let goal ← Elab.Tactic.getMainGoal let goal ← Elab.Tactic.getMainGoal
let .some (beta, _, _) := (← goal.getType).eq? | throwError "Goal is not an Eq" let .some (β, _, _) := (← goal.getType).eq? | throwError "Goal is not an Eq"
-- Create the descendant goals let userName := (← goal.getDecl).userName
let nextGoals ← goal.withContext do let nextGoals ← goal.withContext do
let u ← Meta.mkFreshLevelMVar let u ← Meta.mkFreshLevelMVar
let alpha ← Meta.mkFreshExprMVar (.some $ mkSort u) .natural .anonymous let α ← Meta.mkFreshExprMVar (.some $ mkSort u)
let fType := .forallE .anonymous alpha beta .default .natural (userName := userName ++ `α)
let fType := .forallE .anonymous α β .default
let f₁ ← Meta.mkFreshExprMVar (.some fType) let f₁ ← Meta.mkFreshExprMVar (.some fType)
.synthetic (userName := goal.name ++ `f₁) .synthetic (userName := userName ++ `f₁)
let f₂ ← Meta.mkFreshExprMVar (.some fType) let f₂ ← Meta.mkFreshExprMVar (.some fType)
.synthetic (userName := goal.name ++ `f₂) .synthetic (userName := userName ++ `f₂)
let a₁ ← Meta.mkFreshExprMVar (.some alpha) let a₁ ← Meta.mkFreshExprMVar (.some α)
.synthetic (userName := goal.name ++ `a₁) .synthetic (userName := userName ++ `a₁)
let a₂ ← Meta.mkFreshExprMVar (.some alpha) let a₂ ← Meta.mkFreshExprMVar (.some α)
.synthetic (userName := goal.name ++ `a₂) .synthetic (userName := userName ++ `a₂)
let h₁ ← Meta.mkEq f₁ f₂ let h₁ ← Meta.mkFreshExprMVar (.some $ ← Meta.mkEq f₁ f₂)
let h₂ ← Meta.mkEq a₁ a₂ .synthetic (userName := userName ++ `h₁)
let h₂ ← Meta.mkFreshExprMVar (.some $ ← Meta.mkEq a₁ a₂)
.synthetic (userName := userName ++ `h₂)
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a₁) (.app f₂ a₂)) (← goal.getType) let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a₁) (.app f₂ a₂)) (← goal.getType)
let conduit ← Meta.mkFreshExprMVar conduitType let conduit ← Meta.mkFreshExprMVar conduitType
.synthetic (userName := goal.name ++ `conduit) .synthetic (userName := userName ++ `conduit)
goal.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongr h₁ h₂) goal.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongr h₁ h₂)
return [alpha, f₁, f₂, a₁, a₂, h₁, h₂, conduit] return [α, f₁, f₂, a₁, a₂, h₁, h₂, conduit]
Elab.Tactic.setGoals <| nextGoals.map (·.mvarId!) Elab.Tactic.setGoals <| nextGoals.map (·.mvarId!)
end Pantograph.Tactic end Pantograph.Tactic

View File

@ -27,6 +27,7 @@ def Goal.devolatilize (goal: Goal): Goal :=
name := "", name := "",
} }
deriving instance DecidableEq, Repr for Name
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
@ -65,9 +66,23 @@ def runTermElabMInMeta { α } (termElabM: Lean.Elab.TermElabM α): Lean.MetaM α
def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e
def parseSentence (s: String): MetaM Expr := do
let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv)
(catName := `term)
(input := s)
(fileName := filename) with
| .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}"
runTermElabMInMeta $ Elab.Term.elabTerm (stx := recursor) .none
def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] } let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
return newGoals.goals return newGoals.goals
def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do
let name := (← mvarId.getDecl).userName
let t ← exprToStr (← mvarId.getType)
return (name, t)
end Test end Test

View File

@ -49,6 +49,7 @@ def main (args: List String) := do
("Metavar", Metavar.suite env_default), ("Metavar", Metavar.suite env_default),
("Proofs", Proofs.suite env_default), ("Proofs", Proofs.suite env_default),
("Serial", Serial.suite env_default), ("Serial", Serial.suite env_default),
("Tactic/Congruence", Tactic.Congruence.suite env_default),
("Tactic/Motivated Apply", Tactic.MotivatedApply.suite env_default), ("Tactic/Motivated Apply", Tactic.MotivatedApply.suite env_default),
("Tactic/No Confuse", Tactic.NoConfuse.suite env_default), ("Tactic/No Confuse", Tactic.NoConfuse.suite env_default),
] ]

View File

@ -1,2 +1,3 @@
import Test.Tactic.Congruence
import Test.Tactic.MotivatedApply import Test.Tactic.MotivatedApply
import Test.Tactic.NoConfuse import Test.Tactic.NoConfuse

View File

@ -0,0 +1,36 @@
import LSpec
import Lean
import Test.Common
open Lean
open Pantograph
namespace Pantograph.Test.Tactic.Congruence
def test_congr_arg (env: Environment): IO LSpec.TestSeq :=
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 suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[
("congrArg", test_congr_arg env),
]
end Pantograph.Test.Tactic.Congruence

View File

@ -7,24 +7,11 @@ open Pantograph
namespace Pantograph.Test.Tactic.MotivatedApply namespace Pantograph.Test.Tactic.MotivatedApply
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 := def test_type_extract (env: Environment): IO LSpec.TestSeq :=
runMetaMSeq env do runMetaMSeq env do
let mut tests := LSpec.TestSeq.done let mut tests := LSpec.TestSeq.done
let (recursor, recursorType) ← valueAndType "@Nat.brecOn" 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" = tests := tests ++ LSpec.check "recursorType" ("{motive : Nat → Sort ?u.1} → (t : Nat) → ((t : Nat) → Nat.below t → motive t) → motive t" =
(← exprToStr recursorType)) (← exprToStr recursorType))
let info ← match Tactic.getRecursorInformation recursorType with let info ← match Tactic.getRecursorInformation recursorType with
@ -39,7 +26,7 @@ def test_type_extract (env: Environment): IO LSpec.TestSeq :=
def test_nat_brec_on (env: Environment): IO LSpec.TestSeq := def test_nat_brec_on (env: Environment): IO LSpec.TestSeq :=
let expr := "λ (n t: Nat) => n + 0 = n" let expr := "λ (n t: Nat) => n + 0 = n"
runMetaMSeq env do runMetaMSeq env do
let (expr, exprType) ← valueAndType expr let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
@ -59,7 +46,7 @@ def test_nat_brec_on (env: Environment): IO LSpec.TestSeq :=
"Nat → Prop", "Nat → Prop",
"Nat", "Nat",
"∀ (t : Nat), Nat.below t → ?motive t", "∀ (t : Nat), Nat.below t → ?motive t",
"?motive ?m.69 = (n + 0 = n)", "?motive ?m.67 = (n + 0 = n)",
]) ])
tests := tests ++ test tests := tests ++ test
return tests return tests
@ -67,7 +54,7 @@ def test_nat_brec_on (env: Environment): IO LSpec.TestSeq :=
def test_list_brec_on (env: Environment): IO LSpec.TestSeq := def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l" let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
runMetaMSeq env do runMetaMSeq env do
let (expr, exprType) ← valueAndType expr let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
@ -84,11 +71,11 @@ def test_list_brec_on (env: Environment): IO LSpec.TestSeq :=
let newGoals ← runTacticOnMVar tactic target.mvarId! let newGoals ← runTacticOnMVar tactic target.mvarId!
pure $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = pure $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[ [
"Type ?u.92", "Type ?u.90",
"List ?m.94 → Prop", "List ?m.92 → Prop",
"List ?m.94", "List ?m.92",
"∀ (t : List ?m.94), List.below t → ?motive t", "∀ (t : List ?m.92), List.below t → ?motive t",
"?motive ?m.96 = (l ++ [] = [] ++ l)", "?motive ?m.94 = (l ++ [] = [] ++ l)",
]) ])
tests := tests ++ test tests := tests ++ test
return tests return tests
@ -103,30 +90,31 @@ def test_partial_motive_instantiation (env: Environment): IO LSpec.TestSeq := do
(fileName := filename) with (fileName := filename) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
let (expr, exprType) ← valueAndType expr let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do Meta.lambdaTelescope expr $ λ _ body => do
let mut tests := LSpec.TestSeq.done let mut tests := LSpec.TestSeq.done
-- Apply the tactic -- Apply the tactic
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
let tactic := Tactic.motivatedApply recursor let tactic := Tactic.motivatedApply recursor
let newGoals ← runTacticOnMVar tactic target.mvarId! let newGoals ← runTacticOnMVar tactic target.mvarId!
let majorId := 67
tests := tests ++ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = tests := tests ++ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
[ [
"Nat → Prop", "Nat → Prop",
"Nat", "Nat",
"∀ (t : Nat), Nat.below t → ?motive t", "∀ (t : Nat), Nat.below t → ?motive t",
"?motive ?m.69 = (n + 0 = n)", s!"?motive ?m.{majorId} = (n + 0 = n)",
])) ]))
let [motive, major, step, conduit] := newGoals | panic! "Incorrect goal number" let [motive, major, step, conduit] := newGoals | panic! "Incorrect goal number"
tests := tests ++ (LSpec.check "goal name" (major.name.toString = "_uniq.69")) tests := tests ++ (LSpec.check "goal name" (major.name.toString = s!"_uniq.{majorId}"))
-- Assign motive to `λ x => x + _` -- Assign motive to `λ x => x + _`
let (motive_assign, _) ← valueAndType "λ (x: Nat) => @Nat.add x + 0 = _" let motive_assign ← parseSentence "λ (x: Nat) => @Nat.add x + 0 = _"
motive.assign motive_assign motive.assign motive_assign
let test ← conduit.withContext do let test ← conduit.withContext do
let t := toString (← Meta.ppExpr $ ← conduit.getType) let t := toString (← Meta.ppExpr $ ← conduit.getType)
return LSpec.check "conduit" (t = "(?m.69.add + 0 = ?m.140 ?m.69) = (n + 0 = n)") return LSpec.check "conduit" (t = s!"(?m.{majorId}.add + 0 = ?m.138 ?m.{majorId}) = (n + 0 = n)")
tests := tests ++ test tests := tests ++ test
return tests return tests

View File

@ -7,23 +7,10 @@ open Pantograph
namespace Pantograph.Test.Tactic.NoConfuse namespace Pantograph.Test.Tactic.NoConfuse
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_nat (env: Environment): IO LSpec.TestSeq := def test_nat (env: Environment): IO LSpec.TestSeq :=
let expr := "λ (n: Nat) (h: 0 = n + 1) => False" let expr := "λ (n: Nat) (h: 0 = n + 1) => False"
runMetaMSeq env do runMetaMSeq env do
let (expr, exprType) ← valueAndType expr let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
@ -46,7 +33,7 @@ def test_nat (env: Environment): IO LSpec.TestSeq :=
def test_nat_fail (env: Environment): IO LSpec.TestSeq := def test_nat_fail (env: Environment): IO LSpec.TestSeq :=
let expr := "λ (n: Nat) (h: n = n) => False" let expr := "λ (n: Nat) (h: n = n) => False"
runMetaMSeq env do runMetaMSeq env do
let (expr, _) ← valueAndType expr let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
@ -70,7 +57,7 @@ def test_nat_fail (env: Environment): IO LSpec.TestSeq :=
def test_list (env: Environment): IO LSpec.TestSeq := def test_list (env: Environment): IO LSpec.TestSeq :=
let expr := "λ (l: List Nat) (h: [] = 1 :: l) => False" let expr := "λ (l: List Nat) (h: [] = 1 :: l) => False"
runMetaMSeq env do runMetaMSeq env do
let (expr, exprType) ← valueAndType expr let expr ← parseSentence expr
Meta.lambdaTelescope expr $ λ _ body => do Meta.lambdaTelescope expr $ λ _ body => do
let recursor ← match Parser.runParserCategory let recursor ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
@ -92,9 +79,9 @@ def test_list (env: Environment): IO LSpec.TestSeq :=
def suite (env: Environment): List (String × IO LSpec.TestSeq) := def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[ [
("nat", test_nat env), ("Nat", test_nat env),
("nat_fail", test_nat_fail env), ("Nat fail", test_nat_fail env),
("list", test_list env), ("List", test_list env),
] ]
end Pantograph.Test.Tactic.NoConfuse end Pantograph.Test.Tactic.NoConfuse