Skip to content

Commit 5d84988

Browse files
authored
feat: lake: use system cache for bootstrap (#10621)
This PR alters the Lake directory detection so that the core build (i.e., `bootstrap = true`) is stored in the user cache directory (if available) and never in a toolchain-specific directory. It is also fixes some issues with cache environment configuration discovered along the way.
1 parent 5ede2bf commit 5d84988

File tree

4 files changed

+65
-26
lines changed

4 files changed

+65
-26
lines changed

src/lake/Lake/Config/Env.lean

Lines changed: 58 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -41,19 +41,28 @@ public structure Env where
4141
noCache : Bool
4242
/-- Whether the Lake artifact cache should be enabled by default (i.e., `LAKE_ARTIFACT_CACHE`). -/
4343
enableArtifactCache : Bool
44+
/-- Whether the system cache has been disabled (`LAKE_CACHE_DIR` is set but empty). -/
45+
noSystemCache : Bool := false
4446
/--
45-
The system directory for the Lake cache. Set by `LAKE_CACHE_DIR`.
47+
The directory for the Lake cache. Customized by `LAKE_CACHE_DIR`.
48+
4649
If `none`, no suitable system directory for the cache exists.
4750
-/
48-
lakeCache? : Option Cache
51+
lakeCache? : Option Cache := none
52+
/--
53+
The directory for the Lake cache. Customized by `LAKE_CACHE_DIR`.
54+
55+
Unlike `lakeCache?`, this excludes the toolchain directory from consideration.
56+
57+
If `none`, no suitable system directory for the cache exists.
58+
-/
59+
lakeSystemCache? : Option Cache := none
4960
/-- The authentication key for cache uploads (i.e., `LAKE_CACHE_KEY`). -/
5061
cacheKey? : Option String
5162
/-- The base URL for artifact uploads and downloads from the cache (i.e., `LAKE_CACHE_ARTIFACT_ENDPOINT`). -/
5263
cacheArtifactEndpoint? : Option String
5364
/-- The base URL for revision uploads and downloads from the cache (i.e., `LAKE_CACHE_REVISION_ENDPOINT`). -/
5465
cacheRevisionEndpoint? : Option String
55-
/-- The initial Elan toolchain of the environment (i.e., `ELAN_TOOLCHAIN`). -/
56-
initToolchain : String
5766
/-- The initial Lean library search path of the environment (i.e., `LEAN_PATH`). -/
5867
initLeanPath : SearchPath
5968
/-- The initial Lean source search path of the environment (i.e., `LEAN_SRC_PATH`). -/
@@ -65,7 +74,7 @@ public structure Env where
6574
/--
6675
The preferred toolchain of the environment. May be empty.
6776
68-
Either `initToolchain` or, if none, Lake's `Lean.toolchain`.
77+
Either `ELAN_TOOLCHAIN` or, if none, Lake's `Lean.toolchain`.
6978
-/
7079
toolchain : String
7180
deriving Inhabited
@@ -92,26 +101,45 @@ public def getSystemCacheHome? : BaseIO (Option FilePath) := do
92101
else
93102
return none
94103

104+
def ElanInstall.lakeToolchainCache (elan : ElanInstall) (toolchain : String) : Cache :=
105+
⟨elan.toolchainDir toolchain / "lake" / "cache"
106+
107+
@[inline] def ElanInstall.lakeToolchainCache? (elan : ElanInstall) (toolchain : String) : Option Cache :=
108+
if toolchain.isEmpty then
109+
none
110+
else
111+
some (elan.lakeToolchainCache toolchain)
112+
95113
namespace Env
96114

97115
/-- Compute the Lean toolchain string used by Lake from the process environment. -/
98116
public def computeToolchain : BaseIO String := do
99117
return (← IO.getEnv "ELAN_TOOLCHAIN").getD Lean.toolchain
100118

119+
def computeEnvCache? : BaseIO (Option Cache) := OptionT.run do
120+
let cacheDir ← IO.getEnv "LAKE_CACHE_DIR"
121+
guard cacheDir.isEmpty
122+
return ⟨cacheDir⟩
123+
124+
def computeSystemCache? : BaseIO (Option Cache) := do
125+
return (← getSystemCacheHome?).map (⟨· / "lake"⟩)
126+
127+
def computeToolchainCache? (elan? : Option ElanInstall) (toolchain : String) : Option Cache := do
128+
let elan ← elan?
129+
guard !toolchain.isEmpty
130+
return elan.lakeToolchainCache toolchain
131+
101132
/--
102133
Compute the system cache location from the process environment.
103134
If `none`, no system directory for workspace the cache exists.
104135
-/
105136
public def computeCache? (elan? : Option ElanInstall) (toolchain : String) : BaseIO (Option Cache) := do
106-
if let some cacheDir ← IO.getEnv "LAKE_CACHE_DIR" then
107-
if cacheDir.isEmpty then
108-
return none
109-
else
110-
return some ⟨cacheDir⟩
111-
else if let some elan := elan? then
112-
return some ⟨elan.toolchainDir toolchain / "lake" / "cache"
113-
else if let some cacheHome ← getSystemCacheHome? then
114-
return some ⟨cacheHome / "lake"
137+
if let some cache ← computeEnvCache? then
138+
return some cache
139+
else if let some cache := computeToolchainCache? elan? toolchain then
140+
return some cache
141+
else if let some cache ← computeSystemCache? then
142+
return some cache
115143
else
116144
return none
117145

@@ -124,28 +152,36 @@ public def compute
124152
(noCache : Option Bool := none)
125153
: EIO String Env := do
126154
let reservoirBaseUrl ← getUrlD "RESERVOIR_API_BASE_URL" "https://reservoir.lean-lang.org/api"
127-
let elanToolchain? ← IO.getEnv "ELAN_TOOLCHAIN"
128-
let initToolchain := elanToolchain?.getD ""
129-
let toolchain := elanToolchain?.getD Lean.toolchain
130-
return {
155+
let toolchain ← computeToolchain
156+
addCacheDirs toolchain {
131157
lake, lean, elan?,
132158
pkgUrlMap := ← computePkgUrlMap
133159
reservoirApiUrl := ← getUrlD "RESERVOIR_API_URL" s!"{reservoirBaseUrl}/v1"
134160
noCache := (noCache <|> (← IO.getEnv "LAKE_NO_CACHE").bind envToBool?).getD false
135161
enableArtifactCache := (← IO.getEnv "LAKE_ARTIFACT_CACHE").bind envToBool? |>.getD false
136-
lakeCache? := ← computeCache? elan? toolchain
137162
cacheKey? := (← IO.getEnv "LAKE_CACHE_KEY").map (·.trim)
138163
cacheArtifactEndpoint? := (← IO.getEnv "LAKE_CACHE_ARTIFACT_ENDPOINT").map normalizeUrl
139164
cacheRevisionEndpoint? := (← IO.getEnv "LAKE_CACHE_REVISION_ENDPOINT").map normalizeUrl
140165
githashOverride := (← IO.getEnv "LEAN_GITHASH").getD ""
141166
toolchain
142-
initToolchain
143167
initLeanPath := ← getSearchPath "LEAN_PATH",
144168
initLeanSrcPath := ← getSearchPath "LEAN_SRC_PATH",
145169
initSharedLibPath := ← getSearchPath sharedLibPathEnvVar,
146170
initPath := ← getSearchPath "PATH"
147171
}
148172
where
173+
addCacheDirs toolchain env := do
174+
if let some dir ← IO.getEnv "LAKE_CACHE_DIR" then
175+
if dir.isEmpty then
176+
return {env with noSystemCache := true}
177+
else
178+
return {env with lakeCache? := some ⟨dir⟩, lakeSystemCache? := some ⟨dir⟩}
179+
else if let some cache := computeToolchainCache? elan? toolchain then
180+
return {env with lakeCache? := some cache, lakeSystemCache? := ← computeSystemCache?}
181+
else if let some cache ← computeSystemCache? then
182+
return {env with lakeCache? := some cache, lakeSystemCache? := some cache}
183+
else
184+
return env
149185
computePkgUrlMap := do
150186
let some urlMapStr ← IO.getEnv "LAKE_PKG_URL_MAP" | return {}
151187
match Json.parse urlMapStr |>.bind fromJson? with
@@ -210,10 +246,11 @@ public def sharedLibPath (env : Env) : SearchPath :=
210246
env.lean.sharedLibPath ++ env.initSharedLibPath
211247

212248
/-- Unset toolchain-specific environment variables. -/
213-
public def noToolchainVars : Array (String × Option String) :=
249+
public def noToolchainVars (env : Env) : Array (String × Option String) :=
214250
#[
215251
("ELAN_TOOLCHAIN", none),
216252
("LAKE", none),
253+
("LAKE_CACHE_DIR", if env.noSystemCache then some "" else env.lakeSystemCache?.map (·.dir.toString)),
217254
("LAKE_OVERRIDE_LEAN", none),
218255
("LAKE_HOME", none),
219256
("LEAN", none),

src/lake/Lake/Config/Workspace.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -25,6 +25,10 @@ public structure Workspace : Type where
2525
root : Package
2626
/-- The detected `Lake.Env` of the workspace. -/
2727
lakeEnv : Lake.Env
28+
/-- The Lake cache. -/
29+
lakeCache : Cache :=
30+
if root.bootstrap then lakeEnv.lakeSystemCache?.getD ⟨root.lakeDir / "cache"
31+
else lakeEnv.lakeCache?.getD ⟨root.lakeDir / "cache"
2832
/--
2933
The CLI arguments Lake was run with.
3034
Used by `lake update` to perform a restart of Lake on a toolchain update.
@@ -69,10 +73,6 @@ namespace Workspace
6973
@[inline] public def lakeDir (self : Workspace) : FilePath :=
7074
self.root.lakeDir
7175

72-
/-- The Lake cache. -/
73-
@[inline] public def lakeCache (ws : Workspace) : Cache :=
74-
ws.lakeEnv.lakeCache?.getD ⟨ws.lakeDir⟩
75-
7676
/-- Whether the Lake artifact cache should be enabled by default for packages in the workspace. -/
7777
@[inline] public def enableArtifactCache (ws : Workspace) : Bool :=
7878
ws.lakeEnv.enableArtifactCache

src/lake/Lake/Load/Resolve.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -256,7 +256,7 @@ def Workspace.updateToolchain
256256
let child ← IO.Process.spawn {
257257
cmd := elanInstall.elan.toString
258258
args := #["run", "--install", tc.toString, "lake"] ++ lakeArgs
259-
env := Env.noToolchainVars
259+
env := ws.lakeEnv.noToolchainVars
260260
}
261261
IO.Process.exit (← child.wait).toUInt8
262262
else

src/lake/tests/env/test.sh

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -36,6 +36,8 @@ LAKE_CACHE_DIR= test_run env -d ../../examples/hello env printenv LAKE_CACHE_DIR
3636
# Test that `env` preserves the input environment for certain variables
3737
echo "# TEST: Setting variables for lake env"
3838
test_eq "foo" env env ELAN_TOOLCHAIN=foo $LAKE env printenv ELAN_TOOLCHAIN
39+
test_out "foo" env env -u LAKE_CACHE_DIR ELAN_HOME=/ ELAN_TOOLCHAIN=foo \
40+
$LAKE env printenv LAKE_CACHE_DIR
3941
LAKE_CACHE_DIR=foo test_eq "foo" env printenv LAKE_CACHE_DIR
4042
LEAN_GITHASH=foo test_eq "foo" env printenv LEAN_GITHASH
4143
LEAN_AR=foo test_eq "foo" env printenv LEAN_AR

0 commit comments

Comments
 (0)