diff --git a/Pantograph/Goal.lean b/Pantograph/Goal.lean index 79e3004..e190d5d 100644 --- a/Pantograph/Goal.lean +++ b/Pantograph/Goal.lean @@ -10,8 +10,6 @@ import Lean namespace Pantograph open Lean -def filename: String := "" - /-- Represents an interconnected set of metavariables, or a state in proof search -/ @@ -73,6 +71,8 @@ protected def GoalState.metaContextOfGoal (state: GoalState) (mvarId: MVarId): O return { lctx := mvarDecl.lctx, localInstances := mvarDecl.localInstances } protected def GoalState.metaState (state: GoalState): Meta.State := state.savedState.term.meta.meta +protected def GoalState.coreState (state: GoalState): Core.SavedState := + state.savedState.term.meta.core protected def GoalState.withContext (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do mvarId.withContext m |>.run' (← read) state.metaState @@ -202,16 +202,27 @@ inductive TacticResult where -- The given action cannot be executed in the state | invalidAction (message: String) -/-- Executes a `TacticM` monads on this `GoalState`, collecting the errors as necessary -/ +/-- Executes a `TacticM` monad on this `GoalState`, collecting the errors as necessary -/ protected def GoalState.tryTacticM (state: GoalState) (goal: MVarId) (tacticM: Elab.Tactic.TacticM Unit): Elab.TermElabM TacticResult := do try let nextState ← state.step goal tacticM + + -- Check if error messages have been generated in the core. + let newMessages ← (← Core.getMessageLog).toList.drop state.coreState.messages.toList.length + |>.filterMapM λ m => do + if m.severity == .error then + return .some $ ← m.toString + else + return .none + if ¬ newMessages.isEmpty then + return .failure newMessages.toArray return .success nextState catch exception => return .failure #[← exception.toMessageData.toString] /-- Execute a string tactic on given state. Restores TermElabM -/ +@[export pantograph_goal_state_try_tactic_m] protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: String): Elab.TermElabM TacticResult := do state.restoreElabM @@ -219,7 +230,7 @@ protected def GoalState.tryTactic (state: GoalState) (goal: MVarId) (tactic: Str (env := ← MonadEnv.getEnv) (catName := if state.isConv then `conv else `tactic) (input := tactic) - (fileName := filename) with + (fileName := ← getFileName) with | .ok stx => pure $ stx | .error error => return .parseError error state.tryTacticM goal $ Elab.Tactic.evalTactic tactic @@ -231,7 +242,7 @@ protected def GoalState.tryAssign (state: GoalState) (goal: MVarId) (expr: Strin (env := ← MonadEnv.getEnv) (catName := `term) (input := expr) - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => return .parseError error state.tryTacticM goal $ Tactic.evalAssign expr @@ -245,7 +256,7 @@ protected def GoalState.tryLet (state: GoalState) (goal: MVarId) (binderName: St (env := ← MonadEnv.getEnv) (catName := `term) (input := type) - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => return .parseError error state.tryTacticM goal $ Tactic.evalLet binderName.toName type @@ -332,7 +343,7 @@ protected def GoalState.tryCalc (state: GoalState) (goal: MVarId) (pred: String) (env := state.env) (catName := `term) (input := pred) - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => return .parseError error goal.checkNotAssigned `GoalState.tryCalc @@ -353,7 +364,7 @@ protected def GoalState.tryCalc (state: GoalState) (goal: MVarId) (pred: String) throwErrorAt pred "invalid 'calc' step, relation expected{indentExpr step}" if let some prevRhs := calcPrevRhs? then unless ← Meta.isDefEqGuarded lhs prevRhs do - throwErrorAt pred "invalid 'calc' step, left-hand-side is{indentD m!"{lhs} : {← Meta.inferType lhs}"}\nprevious right-hand-side is{indentD m!"{prevRhs} : {← Meta.inferType prevRhs}"}" -- " + throwErrorAt pred "invalid 'calc' step, left-hand-side is{indentD m!"{lhs} : {← Meta.inferType lhs}"}\nprevious right-hand-side is{indentD m!"{prevRhs} : {← Meta.inferType prevRhs}"}" -- Creates a mvar to represent the proof that the calc tactic solves the -- current branch diff --git a/Test/Common.lean b/Test/Common.lean index 0a0b44c..53adaa0 100644 --- a/Test/Common.lean +++ b/Test/Common.lean @@ -95,19 +95,19 @@ def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq) def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e -def strToTermSyntax [Monad m] [MonadEnv m] (s: String): m Syntax := do +def strToTermSyntax (s: String): CoreM Syntax := do let .ok stx := Parser.runParserCategory (env := ← MonadEnv.getEnv) (catName := `term) (input := s) - (fileName := filename) | panic! s!"Failed to parse {s}" + (fileName := ← getFileName) | panic! s!"Failed to parse {s}" return stx def parseSentence (s: String): Elab.TermElabM Expr := do let stx ← match Parser.runParserCategory (env := ← MonadEnv.getEnv) (catName := `term) (input := s) - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" Elab.Term.elabTerm (stx := stx) .none @@ -123,23 +123,29 @@ def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do -- Monadic testing -abbrev TestT := StateT LSpec.TestSeq +abbrev TestT := StateRefT' IO.RealWorld LSpec.TestSeq -def addTest [Monad m] (test: LSpec.TestSeq) : TestT m Unit := do +section Monadic + +variable [Monad m] [MonadLiftT (ST IO.RealWorld) m] + +def addTest (test: LSpec.TestSeq) : TestT m Unit := do set $ (← get) ++ test -def checkEq [Monad m] [DecidableEq α] (desc : String) (lhs rhs : α) : TestT m Unit := do - addTest $ LSpec.check desc (lhs == rhs) -def checkTrue [Monad m] (desc : String) (flag : Bool) : TestT m Unit := do +def checkEq [DecidableEq α] [Repr α] (desc : String) (lhs rhs : α) : TestT m Unit := do + addTest $ LSpec.check desc (lhs = rhs) +def checkTrue (desc : String) (flag : Bool) : TestT m Unit := do addTest $ LSpec.check desc flag -def fail [Monad m] (desc : String) : TestT m Unit := do +def fail (desc : String) : TestT m Unit := do addTest $ LSpec.check desc false -def runTest [Monad m] (t: TestT m Unit): m LSpec.TestSeq := +def runTest (t: TestT m Unit): m LSpec.TestSeq := Prod.snd <$> t.run LSpec.TestSeq.done -def runTestWithResult { α } [Monad m] (t: TestT m α): m (α × LSpec.TestSeq) := +def runTestWithResult { α } (t: TestT m α): m (α × LSpec.TestSeq) := t.run LSpec.TestSeq.done +end Monadic + def runTestTermElabM (env: Environment) (t: TestT Elab.TermElabM Unit): IO LSpec.TestSeq := runTermElabMSeq env $ runTest t diff --git a/Test/Metavar.lean b/Test/Metavar.lean index 506e963..c6fc4f0 100644 --- a/Test/Metavar.lean +++ b/Test/Metavar.lean @@ -8,10 +8,7 @@ namespace Pantograph.Test.Metavar open Pantograph open Lean -abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM) - -def addTest (test: LSpec.TestSeq): TestM Unit := do - set $ (← get) ++ test +abbrev TestM := TestT $ ReaderT Protocol.Options Elab.TermElabM -- Tests that all delay assigned mvars are instantiated def test_instantiate_mvar: TestM Unit := do @@ -32,8 +29,6 @@ def test_instantiate_mvar: TestM Unit := do "((:c LE.le) (:c Nat) (:c instLENat) ((:c OfNat.ofNat) (:mv _uniq.2) (:lit 2) (:mv _uniq.3)) ((:c OfNat.ofNat) (:mv _uniq.14) (:lit 5) (:mv _uniq.15)))") return () - - def startProof (expr: String): TestM (Option GoalState) := do let env ← Lean.MonadEnv.getEnv let syn? := parseTerm env expr diff --git a/Test/Proofs.lean b/Test/Proofs.lean index 02c1bf6..65f099f 100644 --- a/Test/Proofs.lean +++ b/Test/Proofs.lean @@ -14,10 +14,7 @@ inductive Start where | copy (name: String) -- Start from some name in the environment | expr (expr: String) -- Start from some expression -abbrev TestM := StateRefT LSpec.TestSeq (ReaderT Protocol.Options Elab.TermElabM) - -def addTest (test: LSpec.TestSeq): TestM Unit := do - set $ (← get) ++ test +abbrev TestM := TestT $ ReaderT Protocol.Options $ Elab.TermElabM def startProof (start: Start): TestM (Option GoalState) := do let env ← Lean.MonadEnv.getEnv @@ -704,6 +701,25 @@ def test_nat_zero_add_alt: TestM Unit := do } ]) +def test_composite_tactic_failure: TestM Unit := do + let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)") + let state0 ← match state? with + | .some state => pure state + | .none => do + addTest $ assertUnreachable "Goal could not parse" + return () + + let tactic := "intro p" + let state1 ← match ← state0.tacticOn 0 tactic with + | .success state => pure state + | other => do + addTest $ assertUnreachable $ other.toString + return () + + let tactic := "exact ⟨0, by simp⟩" + let .failure messages ← state1.tacticOn 0 tactic | addTest $ assertUnreachable s!"{tactic} should fail" + checkEq s!"{tactic} fails" messages #[s!"{← getFileName}:0:12: error: unsolved goals\np : Nat → Prop\n⊢ p 0\n"] + def suite (env: Environment): List (String × IO LSpec.TestSeq) := let tests := [ ("identity", test_identity), @@ -716,6 +732,7 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) := ("calc", test_calc), ("Nat.zero_add", test_nat_zero_add), ("Nat.zero_add alt", test_nat_zero_add_alt), + ("composite tactic failure", test_composite_tactic_failure), ] tests.map (fun (name, test) => (name, proofRunner env test)) diff --git a/Test/Tactic/MotivatedApply.lean b/Test/Tactic/MotivatedApply.lean index 4fb05e3..61d7d6c 100644 --- a/Test/Tactic/MotivatedApply.lean +++ b/Test/Tactic/MotivatedApply.lean @@ -28,7 +28,7 @@ def test_nat_brec_on : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "@Nat.brecOn") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" -- Apply the tactic @@ -52,7 +52,7 @@ def test_list_brec_on : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "@List.brecOn") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" -- Apply the tactic @@ -74,7 +74,7 @@ def test_partial_motive_instantiation : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "@Nat.brecOn") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" let expr ← parseSentence expr diff --git a/Test/Tactic/NoConfuse.lean b/Test/Tactic/NoConfuse.lean index ac277e2..93f0606 100644 --- a/Test/Tactic/NoConfuse.lean +++ b/Test/Tactic/NoConfuse.lean @@ -15,7 +15,7 @@ def test_nat : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "h") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" -- Apply the tactic @@ -32,7 +32,7 @@ def test_nat_fail : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "h") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" -- Apply the tactic @@ -52,7 +52,7 @@ def test_list : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "h") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" -- Apply the tactic diff --git a/Test/Tactic/Prograde.lean b/Test/Tactic/Prograde.lean index 132718a..b3347cb 100644 --- a/Test/Tactic/Prograde.lean +++ b/Test/Tactic/Prograde.lean @@ -15,7 +15,7 @@ def test_define : TestT Elab.TermElabM Unit := do (env := ← MonadEnv.getEnv) (catName := `term) (input := "Or.inl h") - (fileName := filename) with + (fileName := ← getFileName) with | .ok syn => pure syn | .error error => throwError "Failed to parse: {error}" -- Apply the tactic