Skip to content

Commit

Permalink
fix: make import resolution case-sensitive on all platforms (#4538)
Browse files Browse the repository at this point in the history
Co-authored-by: Mac Malone <tydeu@hatpress.net>
  • Loading branch information
Kha and tydeu authored Jul 31, 2024
1 parent d19bab0 commit d5a8c96
Show file tree
Hide file tree
Showing 9 changed files with 62 additions and 12 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Elab/Import.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ def parseImports (input : String) (fileName : Option String := none) : IO (Array
def printImports (input : String) (fileName : Option String) : IO Unit := do
let (deps, _, _) ← parseImports input fileName
for dep in deps do
let fname ← findOLean dep.module
let fname ← findOLean (checkExists := false) dep.module
IO.println fname

end Lean.Elab
2 changes: 0 additions & 2 deletions src/Lean/Environment.lean
Original file line number Diff line number Diff line change
Expand Up @@ -806,8 +806,6 @@ partial def importModulesCore (imports : Array Import) : ImportStateM Unit := do
continue
modify fun s => { s with moduleNameSet := s.moduleNameSet.insert i.module }
let mFile ← findOLean i.module
unless (← mFile.pathExists) do
throw <| IO.userError s!"object file '{mFile}' of module {i.module} does not exist"
let (mod, region) ← readModuleData mFile
importModulesCore mod.imports
modify fun s => { s with
Expand Down
53 changes: 44 additions & 9 deletions src/Lean/Util/Path.lean
Original file line number Diff line number Diff line change
Expand Up @@ -40,25 +40,52 @@ where
| Name.anonymous => base
| Name.num _ _ => panic! "ill-formed import"

variable (base : FilePath) (ext : String) in
/--
Checks whether a module of the given name and extension exists in `base`; this uses case-sensitive
path comparisons regardless of underlying file system to ensure the check is consistent across
platforms.
-/
partial def moduleExists : Name → IO Bool := go ext
where go (ext : String)
| .mkStr parent str => do
-- Case-sensitive check for file with extension in top-level call, for directory recursively
let entryName := if ext.isEmpty then str else s!"{str}.{ext}"
unless (← go "" parent) do
return false
return (← (modToFilePath base parent "").readDir).any (·.fileName == entryName)
| .anonymous => base.pathExists
| .num .. => panic! "ill-formed import"

/-- A `.olean' search path. -/
abbrev SearchPath := System.SearchPath

namespace SearchPath

def findRoot (sp : SearchPath) (ext : String) (pkg : String) : IO (Option FilePath) := do
sp.findM? fun p => do
unless (← p.isDir) do -- Lake may pass search roots that do not exist (yet)
return false
if (← (p / pkg).isDir) then
return (← p.readDir).any (·.fileName == pkg)
else
let fileName := s!"{pkg}.{ext}"
return (← p.readDir).any (·.fileName == fileName)

/-- If the package of `mod` can be found in `sp`, return the path with extension
`ext` (`lean` or `olean`) corresponding to `mod`. Otherwise, return `none`. Does
not check whether the returned path exists. -/
def findWithExt (sp : SearchPath) (ext : String) (mod : Name) : IO (Option FilePath) := do
let pkg := mod.getRoot.toString (escape := false)
let root? ← sp.findM? fun p =>
(p / pkg).isDir <||> ((p / pkg).addExtension ext).pathExists
let root? ← findRoot sp ext pkg
return root?.map (modToFilePath · mod ext)

/-- Like `findWithExt`, but ensures the returned path exists. -/
def findModuleWithExt (sp : SearchPath) (ext : String) (mod : Name) : IO (Option FilePath) := do
if let some path ← findWithExt sp ext mod then
if ← path.pathExists then
return some path
let pkg := mod.getRoot.toString (escape := false)
if let some root ← findRoot sp ext pkg then
if (← moduleExists root ext mod) then
return some <| modToFilePath root mod ext
return none

def findAllWithExt (sp : SearchPath) (ext : String) : IO (Array FilePath) := do
Expand Down Expand Up @@ -105,12 +132,20 @@ def initSearchPath (leanSysroot : FilePath) (sp : SearchPath := ∅) : IO Unit :
private def initSearchPathInternal : IO Unit := do
initSearchPath (← getBuildDir)

partial def findOLean (mod : Name) : IO FilePath := do
/--
Returns the path of the .olean file for `mod`. Throws an error if no search path entry for `mod`
could be located, or if `checkExists` is true and the resulting path does not exist.
-/
partial def findOLean (mod : Name) (checkExists := true) : IO FilePath := do
let sp ← searchPathRef.get
if let some fname ← sp.findWithExt "olean" mod then
return fname
let pkg := mod.getRoot.toString (escape := false)
if let some root ← sp.findRoot "olean" pkg then
let path := modToFilePath root mod "olean"
if !checkExists || (← moduleExists root "olean" mod) then
return path
else
throw <| IO.userError s!"object file '{path}' of module {mod} does not exist"
else
let pkg := FilePath.mk <| mod.getRoot.toString (escape := false)
let mut msg := s!"unknown module prefix '{pkg}'
No directory '{pkg}' or file '{pkg}.olean' in the search path entries:
Expand Down
3 changes: 3 additions & 0 deletions src/lake/Lake/Build/Module.lean
Original file line number Diff line number Diff line change
Expand Up @@ -31,6 +31,9 @@ Recursively parse the Lean files of a module and its imports
building an `Array` product of its direct local imports.
-/
def Module.recParseImports (mod : Module) : FetchM (Array Module) := do
-- Make sure we fail reading the file independently of the FS case sensitivity
unless (← Lean.moduleExists mod.lib.srcDir "lean" mod.name) do
error s!"module source file not found: {mod.leanFile}"
let contents ← IO.FS.readFile mod.leanFile
let imports ← Lean.parseImports' contents mod.leanFile.toString
let mods ← imports.foldlM (init := OrdModuleSet.empty) fun set imp =>
Expand Down
1 change: 1 addition & 0 deletions tests/pkg/import-case/.gitignore
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
/.lake
2 changes: 2 additions & 0 deletions tests/pkg/import-case/ImportCase.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
-- should *not* be accepted on any platform
import ImportCase.basic
1 change: 1 addition & 0 deletions tests/pkg/import-case/ImportCase/Basic.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
def hello := "world"
5 changes: 5 additions & 0 deletions tests/pkg/import-case/lakefile.toml
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
name = "«import-case»"
defaultTargets = ["ImportCase"]

[[lean_lib]]
name = "ImportCase"
5 changes: 5 additions & 0 deletions tests/pkg/import-case/test.sh
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
#!/usr/bin/env bash
set -euxo pipefail

rm -rf .lake/build
(lake build 2>&1 && exit 1 || true) | grep --color -F 'module source file not found'

0 comments on commit d5a8c96

Please sign in to comment.