Pantograph/Test/Common.lean

101 lines
3.5 KiB
Plaintext
Raw Normal View History

2023-11-04 15:00:51 -07:00
import Pantograph.Goal
2024-03-28 00:06:35 -07:00
import Pantograph.Library
import Pantograph.Protocol
import Lean
2023-11-04 15:00:51 -07:00
import LSpec
open Lean
namespace Pantograph
2024-06-25 12:13:58 -07:00
deriving instance Repr for Expr
-- Use strict equality check for expressions
--instance : BEq Expr := ⟨Expr.equal⟩
instance (priority := 80) (x y : Expr) : LSpec.Testable (x.equal y) :=
if h : Expr.equal x y then
.isTrue h
else
.isFalse h $ s!"Expected to be equalaaa: '{x.dbgToString}' and '{y.dbgToString}'"
def uniq (n: Nat): Name := .num (.str .anonymous "_uniq") n
2024-04-06 14:07:13 -07:00
-- Auxiliary functions
namespace Protocol
2024-04-15 12:47:02 -07:00
def Goal.devolatilizeVars (goal: Goal): Goal :=
{
goal with
vars := goal.vars.map removeInternalAux,
}
where removeInternalAux (v: Variable): Variable :=
{
v with
name := ""
}
2024-04-15 12:47:02 -07:00
/-- Set internal names to "" -/
def Goal.devolatilize (goal: Goal): Goal :=
{
goal.devolatilizeVars with
name := "",
}
2024-05-20 11:51:35 -07:00
deriving instance DecidableEq, Repr for Name
2023-11-04 15:00:51 -07:00
deriving instance DecidableEq, Repr for Expression
deriving instance DecidableEq, Repr for Variable
deriving instance DecidableEq, Repr for Goal
2024-03-31 16:43:30 -07:00
deriving instance DecidableEq, Repr for ExprEchoResult
deriving instance DecidableEq, Repr for InteractionError
deriving instance DecidableEq, Repr for Option
end Protocol
2023-11-04 15:00:51 -07:00
def TacticResult.toString : TacticResult → String
| .success state => s!".success ({state.goals.length} goals)"
| .failure messages =>
let messages := "\n".intercalate messages.toList
s!".failure {messages}"
| .parseError error => s!".parseError {error}"
| .indexError index => s!".indexError {index}"
2024-04-08 12:26:22 -07:00
| .invalidAction error => s!".invalidAction {error}"
2023-11-04 15:00:51 -07:00
2024-04-06 14:07:13 -07:00
namespace Test
2023-11-04 15:00:51 -07:00
2024-04-06 14:07:13 -07:00
def expectationFailure (desc: String) (error: String): LSpec.TestSeq := LSpec.test desc (LSpec.ExpectationFailure "ok _" error)
def assertUnreachable (message: String): LSpec.TestSeq := LSpec.check message false
2024-01-07 14:14:20 -08:00
2024-04-06 17:45:36 -07:00
def parseFailure (error: String) := expectationFailure "parse" error
def elabFailure (error: String) := expectationFailure "elab" error
2024-03-31 16:43:30 -07:00
def runCoreMSeq (env: Environment) (coreM: CoreM LSpec.TestSeq) (options: Array String := #[]): IO LSpec.TestSeq := do
let coreContext: Core.Context ← createCoreContext options
2024-01-07 14:14:20 -08:00
match ← (coreM.run' coreContext { env := env }).toBaseIO with
| .error exception =>
return LSpec.test "Exception" (s!"internal exception #{← exception.toMessageData.toString}" = "")
| .ok a => return a
def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSeq :=
runCoreMSeq env metaM.run'
def runTermElabMInMeta { α } (termElabM: Lean.Elab.TermElabM α): Lean.MetaM α :=
2024-04-06 17:45:36 -07:00
termElabM.run' (ctx := Pantograph.defaultTermElabMContext)
2024-01-07 14:14:20 -08:00
2024-04-22 00:11:41 -07:00
def exprToStr (e: Expr): Lean.MetaM String := toString <$> Meta.ppExpr e
2024-05-20 11:51:35 -07:00
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
2024-04-22 00:11:41 -07:00
def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }
return newGoals.goals
2024-05-20 11:51:35 -07:00
def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do
let name := (← mvarId.getDecl).userName
let t ← exprToStr (← mvarId.getType)
return (name, t)
2024-04-22 00:11:41 -07:00
2024-04-06 14:07:13 -07:00
end Test
2024-03-28 00:06:35 -07:00
end Pantograph