2024-11-30 23:21:16 -08:00
|
|
|
|
import LSpec
|
|
|
|
|
import Test.Common
|
|
|
|
|
import Lean
|
|
|
|
|
import Pantograph.Library
|
|
|
|
|
|
|
|
|
|
open Lean
|
|
|
|
|
|
|
|
|
|
namespace Pantograph.Test.Serial
|
|
|
|
|
|
2024-12-04 10:44:33 -08:00
|
|
|
|
def tempPath : IO System.FilePath := do
|
|
|
|
|
Prod.snd <$> IO.FS.createTempFile
|
|
|
|
|
|
2024-11-30 23:21:16 -08:00
|
|
|
|
structure MultiState where
|
|
|
|
|
coreContext : Core.Context
|
2024-12-04 10:44:33 -08:00
|
|
|
|
env: Environment
|
|
|
|
|
|
|
|
|
|
abbrev TestM := TestT $ StateRefT MultiState $ IO
|
2024-11-30 23:21:16 -08:00
|
|
|
|
|
2024-12-04 10:44:33 -08:00
|
|
|
|
instance : MonadEnv TestM where
|
|
|
|
|
getEnv := return (← getThe MultiState).env
|
|
|
|
|
modifyEnv f := do modifyThe MultiState fun s => { s with env := f s.env }
|
2024-11-30 23:21:16 -08:00
|
|
|
|
|
2024-12-04 10:44:33 -08:00
|
|
|
|
def runCoreM { α } (state : Core.State) (testCoreM : TestT CoreM α) : TestM (α × Core.State) := do
|
|
|
|
|
let multiState ← getThe MultiState
|
2024-11-30 23:21:16 -08:00
|
|
|
|
let coreM := runTestWithResult testCoreM
|
2024-12-04 10:44:33 -08:00
|
|
|
|
match ← (coreM.run multiState.coreContext state).toBaseIO with
|
|
|
|
|
| .error e => do
|
|
|
|
|
throw $ .userError $ ← e.toMessageData.toString
|
|
|
|
|
| .ok ((a, tests), state') => do
|
2024-11-30 23:21:16 -08:00
|
|
|
|
set $ (← getThe LSpec.TestSeq) ++ tests
|
2024-12-04 10:44:33 -08:00
|
|
|
|
return (a, state')
|
|
|
|
|
|
|
|
|
|
def test_environment_pickling : TestM Unit := do
|
2024-12-05 14:23:55 -08:00
|
|
|
|
let coreSrc : Core.State := { env := ← getEnv }
|
|
|
|
|
let coreDst : Core.State := { env := ← getEnv }
|
2024-12-04 10:44:33 -08:00
|
|
|
|
|
|
|
|
|
let name := `mystery
|
|
|
|
|
let envPicklePath ← tempPath
|
2024-12-05 14:23:55 -08:00
|
|
|
|
let ((), _) ← runCoreM coreSrc do
|
2024-12-04 10:44:33 -08:00
|
|
|
|
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
|
|
|
|
|
let value: Expr := .lam `p (.sort 0) (.lam `h (.bvar 0) (.bvar 0) .default) .default
|
|
|
|
|
let c := Lean.Declaration.defnDecl <| Lean.mkDefinitionValEx
|
|
|
|
|
(name := name)
|
|
|
|
|
(levelParams := [])
|
|
|
|
|
(type := type)
|
|
|
|
|
(value := value)
|
|
|
|
|
(hints := Lean.mkReducibilityHintsRegularEx 1)
|
|
|
|
|
(safety := Lean.DefinitionSafety.safe)
|
|
|
|
|
(all := [])
|
|
|
|
|
let env' ← match (← getEnv).addDecl (← getOptions) c with
|
|
|
|
|
| .error e => do
|
|
|
|
|
let error ← (e.toMessageData (← getOptions)).toString
|
|
|
|
|
throwError error
|
|
|
|
|
| .ok env' => pure env'
|
|
|
|
|
environmentPickle env' envPicklePath
|
|
|
|
|
|
2024-12-05 14:23:55 -08:00
|
|
|
|
let _ ← runCoreM coreDst do
|
2024-12-04 10:44:33 -08:00
|
|
|
|
let (env', _) ← environmentUnpickle envPicklePath
|
|
|
|
|
checkTrue s!"Has symbol {name}" (env'.find? name).isSome
|
|
|
|
|
let anotherName := `mystery2
|
|
|
|
|
checkTrue s!"Doesn't have symbol {anotherName}" (env'.find? anotherName).isNone
|
2024-11-30 23:21:16 -08:00
|
|
|
|
|
2024-12-04 10:44:33 -08:00
|
|
|
|
IO.FS.removeFile envPicklePath
|
2024-11-30 23:21:16 -08:00
|
|
|
|
|
2024-12-05 14:23:55 -08:00
|
|
|
|
def test_goal_state_pickling_simple : TestM Unit := do
|
|
|
|
|
let coreSrc : Core.State := { env := ← getEnv }
|
|
|
|
|
let coreDst : Core.State := { env := ← getEnv }
|
|
|
|
|
let statePath ← tempPath
|
|
|
|
|
|
|
|
|
|
let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default
|
|
|
|
|
let stateGenerate : MetaM GoalState := runTermElabMInMeta do
|
|
|
|
|
GoalState.create type
|
|
|
|
|
|
|
|
|
|
let ((), _) ← runCoreM coreSrc do
|
|
|
|
|
let state ← stateGenerate.run'
|
|
|
|
|
goalStatePickle state statePath
|
|
|
|
|
|
|
|
|
|
let ((), _) ← runCoreM coreDst do
|
|
|
|
|
let (goalState, _) ← goalStateUnpickle statePath (← getEnv)
|
|
|
|
|
let metaM : MetaM (List Expr) := do
|
|
|
|
|
goalState.goals.mapM λ goal => goalState.withContext goal goal.getType
|
|
|
|
|
let types ← metaM.run'
|
|
|
|
|
checkTrue "Goals" $ types[0]!.equal type
|
|
|
|
|
|
|
|
|
|
IO.FS.removeFile statePath
|
|
|
|
|
|
2024-11-30 23:21:16 -08:00
|
|
|
|
structure Test where
|
|
|
|
|
name : String
|
|
|
|
|
routine: TestM Unit
|
|
|
|
|
|
|
|
|
|
protected def Test.run (test: Test) (env: Lean.Environment) : IO LSpec.TestSeq := do
|
|
|
|
|
-- Create the state
|
|
|
|
|
let state : MultiState := {
|
|
|
|
|
coreContext := ← createCoreContext #[],
|
2024-12-04 10:44:33 -08:00
|
|
|
|
env,
|
2024-11-30 23:21:16 -08:00
|
|
|
|
}
|
2024-12-04 10:44:33 -08:00
|
|
|
|
match ← ((runTest $ test.routine).run' state).toBaseIO with
|
2024-11-30 23:21:16 -08:00
|
|
|
|
| .ok e => return e
|
2024-12-04 10:44:33 -08:00
|
|
|
|
| .error e =>
|
2024-12-05 14:23:55 -08:00
|
|
|
|
return LSpec.check s!"Emitted exception: {e.toString}" (e.toString == "")
|
2024-11-30 23:21:16 -08:00
|
|
|
|
|
|
|
|
|
def suite (env : Lean.Environment): List (String × IO LSpec.TestSeq) :=
|
|
|
|
|
let tests: List Test := [
|
2024-12-05 14:23:55 -08:00
|
|
|
|
{ name := "environment_pickling", routine := test_environment_pickling, },
|
|
|
|
|
{ name := "goal_state_pickling_simple", routine := test_goal_state_pickling_simple, },
|
2024-11-30 23:21:16 -08:00
|
|
|
|
]
|
|
|
|
|
tests.map (fun test => (test.name, test.run env))
|
|
|
|
|
|
|
|
|
|
end Pantograph.Test.Serial
|