Skip to content

Commit

Permalink
test: compiling from the interpreter, with common imports
Browse files Browse the repository at this point in the history
hacky fix to windows test

Include test from #2407 as well
  • Loading branch information
kim-em authored and leodemoura committed Aug 16, 2023
1 parent f22695f commit 58d19b8
Show file tree
Hide file tree
Showing 11 changed files with 114 additions and 0 deletions.
3 changes: 3 additions & 0 deletions tests/pkg/frontend/Frontend.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
import Frontend.RegisterOption
import Frontend.Import1
import Frontend.Import2
28 changes: 28 additions & 0 deletions tests/pkg/frontend/Frontend/Compile.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,28 @@
import Lean.Elab.Frontend

open Lean Elab

unsafe def processInput (input : String) (initializers := false) :
IO (Environment × List Message) := do
let fileName := "<input>"
let inputCtx := Parser.mkInputContext input fileName
if initializers then enableInitializersExecution
let (header, parserState, messages) ← Parser.parseHeader inputCtx
let (env, messages) ← processHeader header {} messages inputCtx
let s ← IO.processCommands inputCtx parserState (Command.mkState env messages {}) <&> Frontend.State.commandState
pure (s.env, s.messages.msgs.toList)

open System in
def findLean (mod : Name) : IO FilePath := do
let olean ← findOLean mod
-- Remove a "build/lib/" substring from the path.
let lean := olean.toString.replace (toString (FilePath.mk "build" / "lib") ++ FilePath.pathSeparator.toString) ""
return FilePath.mk lean |>.withExtension "lean"

/-- Read the source code of the named module. -/
def moduleSource (mod : Name) : IO String := do
IO.FS.readFile (← findLean mod)

unsafe def compileModule (mod : Name) (initializers := false) :
IO (Environment × List Message) := do
processInput (← moduleSource mod) initializers
1 change: 1 addition & 0 deletions tests/pkg/frontend/Frontend/Import1.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
import Frontend.RegisterOption
1 change: 1 addition & 0 deletions tests/pkg/frontend/Frontend/Import2.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
import Frontend.RegisterOption
13 changes: 13 additions & 0 deletions tests/pkg/frontend/Frontend/Main.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
import Frontend.Compile

open Lean

unsafe def main (args : List String) : IO UInt32 := do
initSearchPath (← findSysroot)
let mut count : UInt32 := 0
for mod in args do
IO.println s!"Compiling {mod}"
let (_env, msgs) ← compileModule mod.toName true
for m in msgs do IO.println (← m.toString)
if msgs.length > 0 then count := 1
return count
2 changes: 2 additions & 0 deletions tests/pkg/frontend/Frontend/Main_with_Import2.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
import Frontend.Main
import Frontend.Import2
6 changes: 6 additions & 0 deletions tests/pkg/frontend/Frontend/Main_with_Import2_and_eval.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
import Frontend.Main
import Frontend.Import2

#eval show IO _ from do
let r ← main ["Frontend.Import1"]
if r ≠ 0 then throw <| IO.userError "Messages were generated!"
5 changes: 5 additions & 0 deletions tests/pkg/frontend/Frontend/Main_with_eval.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
import Frontend.Main

#eval show IO _ from do
let r ← main ["Frontend.Import1"]
if r ≠ 0 then throw <| IO.userError "Messages were generated!"
6 changes: 6 additions & 0 deletions tests/pkg/frontend/Frontend/RegisterOption.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
import Lean

register_option option : Bool := {
defValue := true
descr := ""
}
17 changes: 17 additions & 0 deletions tests/pkg/frontend/lakefile.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,17 @@
import Lake
open Lake DSL

package frontend

@[default_target]
lean_lib Frontend

lean_exe frontend_with_import2 {
root := `Frontend.Main_with_Import2
supportInterpreter := true
}

lean_exe frontend {
root := `Frontend.Main
supportInterpreter := true
}
32 changes: 32 additions & 0 deletions tests/pkg/frontend/test.sh
Original file line number Diff line number Diff line change
@@ -0,0 +1,32 @@
#!/usr/bin/env bash

rm -rf build
lake build

# Check that we can compile a file which shares with the executable
# a common import using an initializer.
# Here the executable for `frontend_with_import1` imports `Frontend.Import2`.

# This is a minimisation of a situation in which we want to compile a file
# from a project (e.g. Mathlib), so that we can inject another tactic
# implemented in the same project into a goal state from the file.

# This already worked prior to lean4#2423.
lake exe frontend_with_import2 Frontend.Import1 &&

# Check that if we don't import `Frontend.Import2`, we can successfully run
# #eval main ["Frontend.Import1"]
# in the interpreter

# This already worked prior to lean4#2423.
lake build Frontend.Main_with_eval &&

# However if `Main` has imported `Frontend.Import2`, then
# #eval main ["Frontend.Import1"]
# fails to compile `Frontend.Import1` in the interpreter
# prior to lean4#2423.
lake build Frontend.Main_with_Import2_and_eval &&

# Compiling multiple files with a shared import with an initializer
# failed prior to lean4#2423:
lake exe frontend Frontend.Import1 Frontend.Import2

0 comments on commit 58d19b8

Please sign in to comment.