Skip to content

Commit d5f0ce7

Browse files
committed
refactor: lake: all targets produce jobs
1 parent 58c7a4f commit d5f0ce7

16 files changed

+134
-155
lines changed

src/lake/Lake/Build/Facets.lean

+29-29
Original file line numberDiff line numberDiff line change
@@ -53,7 +53,7 @@ The facet which builds all of a module's dependencies
5353
Returns the list of shared libraries to load along with their search path.
5454
-/
5555
abbrev Module.depsFacet := `deps
56-
module_data deps : Job (SearchPath × Array FilePath)
56+
module_data deps : SearchPath × Array FilePath
5757

5858
/--
5959
The core build facet of a Lean file.
@@ -62,58 +62,58 @@ of the module (i.e., `olean`, `ilean`, `c`).
6262
Its trace just includes its dependencies.
6363
-/
6464
abbrev Module.leanArtsFacet := `leanArts
65-
module_data leanArts : Job Unit
65+
module_data leanArts : Unit
6666

6767
/-- The `olean` file produced by `lean`. -/
6868
abbrev Module.oleanFacet := `olean
69-
module_data olean : Job FilePath
69+
module_data olean : FilePath
7070

7171
/-- The `ilean` file produced by `lean`. -/
7272
abbrev Module.ileanFacet := `ilean
73-
module_data ilean : Job FilePath
73+
module_data ilean : FilePath
7474

7575
/-- The C file built from the Lean file via `lean`. -/
7676
abbrev Module.cFacet := `c
77-
module_data c : Job FilePath
77+
module_data c : FilePath
7878

7979
/-- The LLVM BC file built from the Lean file via `lean`. -/
8080
abbrev Module.bcFacet := `bc
81-
module_data bc : Job FilePath
81+
module_data bc : FilePath
8282

8383
/--
8484
The object file `.c.o` built from `c`.
8585
On Windows, this is `c.o.noexport`, on other systems it is `c.o.export`).
8686
-/
8787
abbrev Module.coFacet := `c.o
88-
module_data c.o : Job FilePath
88+
module_data c.o : FilePath
8989

9090
/-- The object file `.c.o.export` built from `c` (with `-DLEAN_EXPORTING`). -/
9191
abbrev Module.coExportFacet := `c.o.export
92-
module_data c.o.export : Job FilePath
92+
module_data c.o.export : FilePath
9393

9494
/-- The object file `.c.o.noexport` built from `c` (without `-DLEAN_EXPORTING`). -/
9595
abbrev Module.coNoExportFacet := `c.o.noexport
96-
module_data c.o.noexport : Job FilePath
96+
module_data c.o.noexport : FilePath
9797

9898
/-- The object file `.bc.o` built from `bc`. -/
9999
abbrev Module.bcoFacet := `bc.o
100-
module_data bc.o : Job FilePath
100+
module_data bc.o : FilePath
101101

102102
/--
103103
The object file built from `c`/`bc`.
104104
On Windows with the C backend, no Lean symbols are exported.
105105
On every other configuration, symbols are exported.
106106
-/
107107
abbrev Module.oFacet := `o
108-
module_data o : Job FilePath
108+
module_data o : FilePath
109109

110110
/-- The object file built from `c`/`bc` (with Lean symbols exported). -/
111111
abbrev Module.oExportFacet := `o.export
112-
module_data o.export : Job FilePath
112+
module_data o.export : FilePath
113113

114114
/-- The object file built from `c`/`bc` (without Lean symbols exported). -/
115115
abbrev Module.oNoExportFacet := `o.noexport
116-
module_data o.noexport : Job FilePath
116+
module_data o.noexport : FilePath
117117

118118

119119
/-! ## Package Facets -/
@@ -123,35 +123,35 @@ A package's *optional* cached build archive (e.g., from Reservoir or GitHub).
123123
Will NOT cause the whole build to fail if the archive cannot be fetched.
124124
-/
125125
abbrev Package.optBuildCacheFacet := `optCache
126-
package_data optCache : Job Bool
126+
package_data optCache : Bool
127127

128128
/--
129129
A package's cached build archive (e.g., from Reservoir or GitHub).
130130
Will cause the whole build to fail if the archive cannot be fetched.
131131
-/
132132
abbrev Package.buildCacheFacet := `cache
133-
package_data cache : Job Unit
133+
package_data cache : Unit
134134

135135
/--
136136
A package's *optional* build archive from Reservoir.
137137
Will NOT cause the whole build to fail if the barrel cannot be fetched.
138138
-/
139139
abbrev Package.optReservoirBarrelFacet := `optBarrel
140-
package_data optBarrel : Job Bool
140+
package_data optBarrel : Bool
141141

142142
/--
143143
A package's Reservoir build archive from Reservoir.
144144
Will cause the whole build to fail if the barrel cannot be fetched.
145145
-/
146146
abbrev Package.reservoirBarrelFacet := `barrel
147-
package_data barrel : Job Unit
147+
package_data barrel : Unit
148148

149149
/--
150150
A package's *optional* build archive from a GitHub release.
151151
Will NOT cause the whole build to fail if the release cannot be fetched.
152152
-/
153153
abbrev Package.optGitHubReleaseFacet := `optRelease
154-
package_data optRelease : Job Bool
154+
package_data optRelease : Bool
155155

156156
@[deprecated optGitHubReleaseFacet (since := "2024-09-27")]
157157
abbrev Package.optReleaseFacet := optGitHubReleaseFacet
@@ -161,49 +161,49 @@ A package's build archive from a GitHub release.
161161
Will cause the whole build to fail if the release cannot be fetched.
162162
-/
163163
abbrev Package.gitHubReleaseFacet := `release
164-
package_data release : Job Unit
164+
package_data release : Unit
165165

166166
@[deprecated gitHubReleaseFacet (since := "2024-09-27")]
167167
abbrev Package.releaseFacet := gitHubReleaseFacet
168168

169169
/-- A package's `extraDepTargets` mixed with its transitive dependencies'. -/
170170
abbrev Package.extraDepFacet := `extraDep
171-
package_data extraDep : Job Unit
171+
package_data extraDep : Unit
172172

173173
/-! ## Target Facets -/
174174

175175
/-- A Lean library's Lean artifacts (i.e., `olean`, `ilean`, `c`). -/
176176
abbrev LeanLib.leanArtsFacet := `leanArts
177-
library_data leanArts : Job Unit
177+
library_data leanArts : Unit
178178

179179
/-- A Lean library's static artifact. -/
180180
abbrev LeanLib.staticFacet := `static
181-
library_data static : Job FilePath
181+
library_data static : FilePath
182182

183183
/-- A Lean library's static artifact (with exported symbols). -/
184184
abbrev LeanLib.staticExportFacet := `static.export
185-
library_data static.export : Job FilePath
185+
library_data static.export : FilePath
186186

187187
/-- A Lean library's shared artifact. -/
188188
abbrev LeanLib.sharedFacet := `shared
189-
library_data shared : Job FilePath
189+
library_data shared : FilePath
190190

191191
/-- A Lean library's `extraDepTargets` mixed with its package's. -/
192192
abbrev LeanLib.extraDepFacet := `extraDep
193-
library_data extraDep : Job Unit
193+
library_data extraDep : Unit
194194

195195
/-- A Lean binary executable. -/
196196
abbrev LeanExe.exeFacet := `leanExe
197-
target_data leanExe : Job FilePath
197+
target_data leanExe : FilePath
198198

199199
/-- A external library's static binary. -/
200200
abbrev ExternLib.staticFacet := `externLib.static
201-
target_data externLib.static : Job FilePath
201+
target_data externLib.static : FilePath
202202

203203
/-- A external library's shared binary. -/
204204
abbrev ExternLib.sharedFacet := `externLib.shared
205-
target_data externLib.shared : Job FilePath
205+
target_data externLib.shared : FilePath
206206

207207
/-- A external library's dynlib. -/
208208
abbrev ExternLib.dynlibFacet := `externLib.dynlib
209-
target_data externLib.dynlib : Job Dynlib
209+
target_data externLib.dynlib : Dynlib

src/lake/Lake/Build/Fetch.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -59,8 +59,8 @@ abbrev RecBuildM := RecBuildT LogIO
5959

6060
/-- A build function for any element of the Lake build index. -/
6161
abbrev IndexBuildFn (m : TypeType v) :=
62-
-- `DFetchFn BuildInfo (BuildData ·.key) m` with less imports
63-
(info : BuildInfo) → m (BuildData info.key)
62+
-- `DFetchFn BuildInfo (Job <| BuildData ·.key) m` with less imports
63+
(info : BuildInfo) → m (Job (BuildData info.key))
6464

6565
/-- A transformer to equip a monad with a build function for the Lake index. -/
6666
abbrev IndexT (m : TypeType v) := EquipT (IndexBuildFn RecBuildM) m
@@ -78,7 +78,7 @@ abbrev FetchM := FetchT LogIO
7878
@[deprecated FetchM (since := "2024-04-30")] abbrev BuildM := BuildT LogIO
7979

8080
/-- Fetch the result associated with the info using the Lake build index. -/
81-
@[inline] def BuildInfo.fetch (self : BuildInfo) [FamilyOut BuildData self.key α] : FetchM α :=
81+
@[inline] def BuildInfo.fetch (self : BuildInfo) [FamilyOut BuildData self.key α] : FetchM (Job α) :=
8282
fun build => cast (by simp) <| build self
8383

8484
export BuildInfo (fetch)

src/lake/Lake/Build/Index.lean

+6-5
Original file line numberDiff line numberDiff line change
@@ -26,13 +26,13 @@ dynamically-typed equivalent.
2626
-/
2727
@[macro_inline] def mkTargetFacetBuild
2828
(facet : Name) (build : FetchM (Job α))
29-
[h : FamilyOut TargetData facet (Job α)]
30-
: FetchM (TargetData facet) :=
29+
[h : FamilyOut TargetData facet α]
30+
: FetchM (Job (TargetData facet)) :=
3131
cast (by rw [← h.family_key_eq_type]) build
3232

3333
def ExternLib.recBuildStatic (lib : ExternLib) : FetchM (Job FilePath) :=
3434
withRegisterJob s!"{lib.staticTargetName.toString}:static" do
35-
lib.config.getJob <$> fetch (lib.pkg.target lib.staticTargetName)
35+
lib.config.getPath <$> fetch (lib.pkg.target lib.staticTargetName)
3636

3737
def ExternLib.recBuildShared (lib : ExternLib) : FetchM (Job FilePath) :=
3838
withRegisterJob s!"{lib.staticTargetName.toString}:shared" do
@@ -47,7 +47,7 @@ def ExternLib.recComputeDynlib (lib : ExternLib) : FetchM (Job Dynlib) := do
4747
-/
4848

4949
/-- Recursive build function for anything in the Lake build index. -/
50-
def recBuildWithIndex : (info : BuildInfo) → FetchM (BuildData info.key)
50+
def recBuildWithIndex : (info : BuildInfo) → FetchM (Job (BuildData info.key))
5151
| .moduleFacet mod facet => do
5252
if let some config := (← getWorkspace).findModuleFacetConfig? facet then
5353
config.build mod
@@ -82,4 +82,5 @@ Run a recursive Lake build using the Lake build index
8282
and a topological / suspending scheduler.
8383
-/
8484
def FetchM.run (x : FetchM α) : RecBuildM α :=
85-
x (inline <| recFetchMemoize BuildInfo.key recBuildWithIndex)
85+
x <| inline <|
86+
recFetchMemoize (β := (Job <| BuildData ·)) BuildInfo.key recBuildWithIndex

src/lake/Lake/Build/Info.lean

+6-6
Original file line numberDiff line numberDiff line change
@@ -125,27 +125,27 @@ definitions.
125125

126126
/-- The direct local imports of the Lean module. -/
127127
abbrev Module.importsFacet := `lean.imports
128-
module_data lean.imports : Job (Array Module)
128+
module_data lean.imports : Array Module
129129

130130
/-- The transitive local imports of the Lean module. -/
131131
abbrev Module.transImportsFacet := `lean.transImports
132-
module_data lean.transImports : Job (Array Module)
132+
module_data lean.transImports : Array Module
133133

134134
/-- The transitive local imports of the Lean module. -/
135135
abbrev Module.precompileImportsFacet := `lean.precompileImports
136-
module_data lean.precompileImports : Job (Array Module)
136+
module_data lean.precompileImports : Array Module
137137

138138
/-- Shared library for `--load-dynlib`. -/
139139
abbrev Module.dynlibFacet := `dynlib
140-
module_data dynlib : Job Dynlib
140+
module_data dynlib : Dynlib
141141

142142
/-- A Lean library's Lean modules. -/
143143
abbrev LeanLib.modulesFacet := `modules
144-
library_data modules : Job (Array Module)
144+
library_data modules : Array Module
145145

146146
/-- The package's complete array of transitive dependencies. -/
147147
abbrev Package.depsFacet := `deps
148-
package_data deps : Job (Array Package)
148+
package_data deps : Array Package
149149

150150

151151
/-!

src/lake/Lake/Build/Package.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -147,7 +147,7 @@ def Package.fetchBuildArchive
147147
private def Package.mkOptBuildArchiveFacetConfig
148148
{facet : Name} (archiveFile : Package → FilePath)
149149
(getUrl : Package → JobM String) (headers : Array String := #[])
150-
[FamilyDef PackageData facet (Job Bool)]
150+
[FamilyDef PackageData facet Bool]
151151
: PackageFacetConfig facet := mkFacetJobConfig fun pkg =>
152152
withRegisterJob s!"{pkg.name}:{facet}" (optional := true) <| Job.async do
153153
try
@@ -161,8 +161,8 @@ private def Package.mkOptBuildArchiveFacetConfig
161161
@[inline]
162162
private def Package.mkBuildArchiveFacetConfig
163163
{facet : Name} (optFacet : Name) (what : String)
164-
[FamilyDef PackageData facet (Job Unit)]
165-
[FamilyDef PackageData optFacet (Job Bool)]
164+
[FamilyDef PackageData facet Unit]
165+
[FamilyDef PackageData optFacet Bool]
166166
: PackageFacetConfig facet :=
167167
mkFacetJobConfig fun pkg =>
168168
withRegisterJob s!"{pkg.name}:{facet}" do

src/lake/Lake/Build/Store.lean

+22-20
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,7 @@ Authors: Mac Malone
55
-/
66
prelude
77
import Lake.Build.Data
8+
import Lake.Build.Job.Basic
89
import Lake.Util.StoreInsts
910

1011
/-!
@@ -19,23 +20,21 @@ namespace Lake
1920
open Lean (Name NameMap)
2021

2122
/-- A monad equipped with a Lake build store. -/
22-
abbrev MonadBuildStore (m) := MonadDStore BuildKey BuildData m
23+
abbrev MonadBuildStore (m) := MonadDStore BuildKey (Job <| BuildData ·) m
2324

2425
/-- The type of the Lake build store. -/
2526
abbrev BuildStore :=
26-
DRBMap BuildKey BuildData BuildKey.quickCmp
27+
DRBMap BuildKey (Job <| BuildData ·) BuildKey.quickCmp
2728

2829
@[inline] def BuildStore.empty : BuildStore := DRBMap.empty
2930

3031
namespace BuildStore
3132

32-
-- Linter reports false positives on the `v` variables below
33-
set_option linter.unusedVariables false
34-
3533
/-- Derive an array of built module facets from the store. -/
36-
def collectModuleFacetArray (self : BuildStore)
37-
(facet : Name) [FamilyOut ModuleData facet α] : Array α := Id.run do
38-
let mut res : Array α := #[]
34+
def collectModuleFacetArray
35+
(self : BuildStore) (facet : Name) [FamilyOut ModuleData facet α]
36+
: Array (Job α) := Id.run do
37+
let mut res : Array (Job α) := #[]
3938
for ⟨k, v⟩ in self do
4039
match k with
4140
| .moduleFacet m f =>
@@ -46,9 +45,10 @@ def collectModuleFacetArray (self : BuildStore)
4645
return res
4746

4847
/-- Derive a map of module names to built facets from the store. -/
49-
def collectModuleFacetMap (self : BuildStore)
50-
(facet : Name) [FamilyOut ModuleData facet α] : NameMap α := Id.run do
51-
let mut res := Lean.mkNameMap α
48+
def collectModuleFacetMap
49+
(self : BuildStore) (facet : Name) [FamilyOut ModuleData facet α]
50+
: NameMap (Job α) := Id.run do
51+
let mut res := Lean.mkNameMap (Job α)
5252
for ⟨k, v⟩ in self do
5353
match k with
5454
| .moduleFacet m f =>
@@ -59,9 +59,10 @@ def collectModuleFacetMap (self : BuildStore)
5959
return res
6060

6161
/-- Derive an array of built package facets from the store. -/
62-
def collectPackageFacetArray (self : BuildStore)
63-
(facet : Name) [FamilyOut PackageData facet α] : Array α := Id.run do
64-
let mut res : Array α := #[]
62+
def collectPackageFacetArray
63+
(self : BuildStore) (facet : Name) [FamilyOut PackageData facet α]
64+
: Array (Job α) := Id.run do
65+
let mut res : Array (Job α) := #[]
6566
for ⟨k, v⟩ in self do
6667
match k with
6768
| .packageFacet _ f =>
@@ -72,9 +73,10 @@ def collectPackageFacetArray (self : BuildStore)
7273
return res
7374

7475
/-- Derive an array of built target facets from the store. -/
75-
def collectTargetFacetArray (self : BuildStore)
76-
(facet : Name) [FamilyOut TargetData facet α] : Array α := Id.run do
77-
let mut res : Array α := #[]
76+
def collectTargetFacetArray
77+
(self : BuildStore) (facet : Name) [FamilyOut TargetData facet α]
78+
: Array (Job α) := Id.run do
79+
let mut res : Array (Job α) := #[]
7880
for ⟨k, v⟩ in self do
7981
match k with
8082
| .targetFacet _ _ f =>
@@ -85,6 +87,6 @@ def collectTargetFacetArray (self : BuildStore)
8587
return res
8688

8789
/-- Derive an array of built external shared libraries from the store. -/
88-
def collectSharedExternLibs (self : BuildStore)
89-
[FamilyOut TargetData `externLib.shared α] : Array α :=
90-
self.collectTargetFacetArray `externLib.shared
90+
def collectSharedExternLibs
91+
(self : BuildStore) [FamilyOut TargetData `externLib.shared α]
92+
: Array (Job α) := self.collectTargetFacetArray `externLib.shared

0 commit comments

Comments
 (0)