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
|
2024-06-25 13:18:31 -07:00
|
|
|
|
import Pantograph.Condensed
|
2024-03-28 18:56:42 -07:00
|
|
|
|
import Lean
|
2023-11-04 15:00:51 -07:00
|
|
|
|
import LSpec
|
2023-10-25 22:18:59 -07:00
|
|
|
|
|
2024-03-28 18:56:42 -07:00
|
|
|
|
open Lean
|
|
|
|
|
|
2023-10-25 22:18:59 -07:00
|
|
|
|
namespace Pantograph
|
|
|
|
|
|
2024-06-25 08:03:08 -07:00
|
|
|
|
deriving instance Repr for Expr
|
|
|
|
|
-- Use strict equality check for expressions
|
2024-06-25 13:18:31 -07:00
|
|
|
|
instance : BEq Expr := ⟨Expr.equal⟩
|
2024-06-25 08:03:08 -07:00
|
|
|
|
|
|
|
|
|
def uniq (n: Nat): Name := .num (.str .anonymous "_uniq") n
|
|
|
|
|
|
2024-04-06 14:07:13 -07:00
|
|
|
|
-- Auxiliary functions
|
2023-10-25 22:18:59 -07:00
|
|
|
|
namespace Protocol
|
2024-04-15 12:47:02 -07:00
|
|
|
|
def Goal.devolatilizeVars (goal: Goal): Goal :=
|
2023-10-25 22:18:59 -07:00
|
|
|
|
{
|
|
|
|
|
goal with
|
|
|
|
|
vars := goal.vars.map removeInternalAux,
|
2024-06-25 13:18:31 -07:00
|
|
|
|
|
2023-10-25 22:18:59 -07:00
|
|
|
|
}
|
|
|
|
|
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
|
2023-10-25 22:18:59 -07:00
|
|
|
|
end Protocol
|
|
|
|
|
|
2024-06-25 13:18:31 -07:00
|
|
|
|
namespace Condensed
|
|
|
|
|
|
|
|
|
|
deriving instance BEq, Repr for LocalDecl
|
|
|
|
|
deriving instance BEq, Repr for Goal
|
|
|
|
|
|
|
|
|
|
protected def LocalDecl.devolatilize (decl: LocalDecl): LocalDecl :=
|
|
|
|
|
{
|
|
|
|
|
decl with fvarId := { name := .anonymous }
|
|
|
|
|
}
|
|
|
|
|
protected def Goal.devolatilize (goal: Goal): Goal :=
|
|
|
|
|
{
|
|
|
|
|
goal with
|
|
|
|
|
mvarId := { name := .anonymous },
|
|
|
|
|
context := goal.context.map LocalDecl.devolatilize
|
|
|
|
|
}
|
|
|
|
|
|
|
|
|
|
end Condensed
|
|
|
|
|
|
2024-09-06 21:07:12 -07:00
|
|
|
|
def GoalState.get! (state: GoalState) (i: Nat): MVarId := state.goals.get! i
|
|
|
|
|
def GoalState.tacticOn (state: GoalState) (goalId: Nat) (tactic: String) := state.tryTactic (state.goals.get! goalId) tactic
|
|
|
|
|
|
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}"
|
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-08-15 22:45:43 -07:00
|
|
|
|
termElabM.run' (ctx := Condensed.elabContext)
|
2024-06-27 11:34:21 -07:00
|
|
|
|
def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq): IO LSpec.TestSeq :=
|
2024-08-15 22:45:43 -07:00
|
|
|
|
runMetaMSeq env $ termElabM.run' (ctx := Condensed.elabContext)
|
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-09-03 14:15:52 -07:00
|
|
|
|
def strToTermSyntax [Monad m] [MonadEnv m] (s: String): m Syntax := do
|
|
|
|
|
let .ok stx := Parser.runParserCategory
|
|
|
|
|
(env := ← MonadEnv.getEnv)
|
|
|
|
|
(catName := `term)
|
|
|
|
|
(input := s)
|
|
|
|
|
(fileName := filename) | panic! s!"Failed to parse {s}"
|
|
|
|
|
return stx
|
2024-06-27 11:34:21 -07:00
|
|
|
|
def parseSentence (s: String): Elab.TermElabM Expr := do
|
2024-09-03 14:15:52 -07:00
|
|
|
|
let stx ← match Parser.runParserCategory
|
2024-05-20 11:51:35 -07:00
|
|
|
|
(env := ← MonadEnv.getEnv)
|
|
|
|
|
(catName := `term)
|
|
|
|
|
(input := s)
|
|
|
|
|
(fileName := filename) with
|
|
|
|
|
| .ok syn => pure syn
|
|
|
|
|
| .error error => throwError "Failed to parse: {error}"
|
2024-09-03 14:15:52 -07:00
|
|
|
|
Elab.Term.elabTerm (stx := stx) .none
|
2024-05-20 11:51:35 -07:00
|
|
|
|
|
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-06-27 11:34:21 -07:00
|
|
|
|
|
|
|
|
|
-- 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 },
|
|
|
|
|
})).toArray
|
|
|
|
|
}
|
|
|
|
|
|
2024-04-06 14:07:13 -07:00
|
|
|
|
end Test
|
2024-03-28 00:06:35 -07:00
|
|
|
|
|
2023-10-25 22:18:59 -07:00
|
|
|
|
end Pantograph
|