Skip to content

Commit

Permalink
change them elsewhere too for good measure, and add a tiny test
Browse files Browse the repository at this point in the history
  • Loading branch information
eric-wieser committed Nov 29, 2023
1 parent 54f01b2 commit 1ec1e0b
Show file tree
Hide file tree
Showing 5 changed files with 13 additions and 10 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Server/Watchdog.lean
Original file line number Diff line number Diff line change
Expand Up @@ -677,7 +677,7 @@ def initAndRunWatchdogAux : ServerM Unit := do
def findWorkerPath : IO System.FilePath := do
let mut workerPath ← IO.appPath
if let some path := (←IO.getEnv "LEAN_SYSROOT") then
workerPath := System.FilePath.mk path / "bin" / "lean" |>.withExtension System.FilePath.exeExtension
workerPath := System.FilePath.mk path / "bin" / "lean" |>.addExtension System.FilePath.exeExtension
if let some path := (←IO.getEnv "LEAN_WORKER_PATH") then
workerPath := System.FilePath.mk path
return workerPath
Expand Down
16 changes: 8 additions & 8 deletions src/lake/Lake/Config/InstallPath.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,7 +12,7 @@ namespace Lake

/-- Standard path of `elan` in a Elan installation. -/
def elanExe (home : FilePath) :=
home / "bin" / "elan" |>.withExtension FilePath.exeExtension
home / "bin" / "elan" |>.addExtension FilePath.exeExtension

/-- Information about the local Elan setup. -/
structure ElanInstall where
Expand All @@ -24,19 +24,19 @@ structure ElanInstall where

/-- Standard path of `lean` in a Lean installation. -/
def leanExe (sysroot : FilePath) :=
sysroot / "bin" / "lean" |>.withExtension FilePath.exeExtension
sysroot / "bin" / "lean" |>.addExtension FilePath.exeExtension

/-- Standard path of `leanc` in a Lean installation. -/
def leancExe (sysroot : FilePath) :=
sysroot / "bin" / "leanc" |>.withExtension FilePath.exeExtension
sysroot / "bin" / "leanc" |>.addExtension FilePath.exeExtension

/-- Standard path of `llvm-ar` in a Lean installation. -/
def leanArExe (sysroot : FilePath) :=
sysroot / "bin" / "llvm-ar" |>.withExtension FilePath.exeExtension
sysroot / "bin" / "llvm-ar" |>.addExtension FilePath.exeExtension

/-- Standard path of `clang` in a Lean installation. -/
def leanCcExe (sysroot : FilePath) :=
sysroot / "bin" / "clang" |>.withExtension FilePath.exeExtension
sysroot / "bin" / "clang" |>.addExtension FilePath.exeExtension

/-- Standard path of `libleanshared` in a Lean installation. -/
def leanSharedLib (sysroot : FilePath) :=
Expand All @@ -45,7 +45,7 @@ def leanSharedLib (sysroot : FilePath) :=
sysroot / "bin"
else
sysroot / "lib" / "lean"
dir / "libleanshared" |>.withExtension sharedLibExt
dir / "libleanshared" |>.addExtension sharedLibExt

/-- Path information about the local Lean installation. -/
structure LeanInstall where
Expand Down Expand Up @@ -80,7 +80,7 @@ def LeanInstall.leanCc? (self : LeanInstall) : Option String :=

/-- Standard path of `lake` in a Lake installation. -/
def lakeExe (buildHome : FilePath) :=
buildHome / "bin" / "lake" |>.withExtension FilePath.exeExtension
buildHome / "bin" / "lake" |>.addExtension FilePath.exeExtension

/-- Path information about the local Lake installation. -/
structure LakeInstall where
Expand Down Expand Up @@ -203,7 +203,7 @@ try to return their joint home by assuming they are both located at `<home>/bin`
def findLakeLeanJointHome? : BaseIO (Option FilePath) := do
if let .ok appPath ← IO.appPath.toBaseIO then
if let some appDir := appPath.parent then
let leanExe := appDir / "lean" |>.withExtension FilePath.exeExtension
let leanExe := appDir / "lean" |>.addExtension FilePath.exeExtension
if (← leanExe.pathExists) then
return appDir.parent
return none
Expand Down
2 changes: 1 addition & 1 deletion src/lake/Lake/Config/LeanExe.lean
Original file line number Diff line number Diff line change
Expand Up @@ -61,7 +61,7 @@ The file name of binary executable
(i.e., `exeName` plus the platform's `exeExtension`).
-/
@[inline] def fileName (self : LeanExe) : FilePath :=
FilePath.withExtension self.config.exeName FilePath.exeExtension
FilePath.addExtension self.config.exeName FilePath.exeExtension

/-- The path to the executable in the package's `binDir`. -/
@[inline] def file (self : LeanExe) : FilePath :=
Expand Down
2 changes: 2 additions & 0 deletions tests/lean/filePath.lean
Original file line number Diff line number Diff line change
Expand Up @@ -30,3 +30,5 @@ def norm (f : FilePath) : String :=
#eval norm <| FilePath.withExtension "a/b.tar.gz" "xz"
#eval norm <| FilePath.withExtension "a/b.tar.gz" ""
#eval norm <| FilePath.withExtension "a/b" "tar.gz"

#eval norm <| FilePath.addExtension "a/b.tar.gz" "bak"
1 change: 1 addition & 0 deletions tests/lean/filePath.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -17,3 +17,4 @@ none
"a/b.tar.xz"
"a/b.tar"
"a/b.tar.gz"
"a/b.tar.xz.bak"

0 comments on commit 1ec1e0b

Please sign in to comment.