Skip to content

Commit

Permalink
refactor: lake: use StateRefT for BuildStore
Browse files Browse the repository at this point in the history
  • Loading branch information
tydeu committed Dec 2, 2024
1 parent 27df5e9 commit 79d8c56
Show file tree
Hide file tree
Showing 2 changed files with 25 additions and 9 deletions.
18 changes: 9 additions & 9 deletions src/lake/Lake/Build/Fetch.lean
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@ namespace Lake

/-- A recursive build of a Lake build store that may encounter a cycle. -/
abbrev RecBuildM :=
CallStackT BuildKey <| BuildT <| ELogT <| StateT BuildStore <| BaseIO
CallStackT BuildKey <| BuildT <| ELogT <| StateRefT BuildStore BaseIO

instance : MonadLift LogIO RecBuildM := ⟨ELogT.takeAndRun⟩

Expand All @@ -31,18 +31,18 @@ Run a `JobM` action in `RecBuildM` (and thus `FetchM`).
Generally, this should not be done, and instead a job action
should be run asynchronously in a Job (e.g., via `Job.async`).
-/
@[inline] def RecBuildM.runJobM (x : JobM α) : RecBuildM α := fun _ ctx log store => do
@[inline] def RecBuildM.runJobM (x : JobM α) : RecBuildM α := fun _ ctx log => do
match (← x ctx {log}) with
| .ok a s => return (.ok a s.log, store)
| .error e s => return (.error e s.log, store)
| .ok a s => return .ok a s.log
| .error e s => return .error e s.log

instance : MonadLift JobM RecBuildM := ⟨RecBuildM.runJobM⟩

/-- Run a recursive build. -/
@[inline] def RecBuildM.run
(stack : CallStack BuildKey) (store : BuildStore) (build : RecBuildM α)
: CoreBuildM (α × BuildStore) := fun ctx log => do
match (← build stack ctx log store) with
match (← (build stack ctx log).run store) with
| (.ok a log, store) => return .ok (a, store) log
| (.error e log, _) => return .error e log

Expand Down Expand Up @@ -91,14 +91,14 @@ def ensureJob (x : FetchM (Job α))
: FetchM (Job α) := fun fetch stack ctx log store => do
let iniPos := log.endPos
match (← (withLoggedIO x) fetch stack ctx log store) with
| (.ok job log, store) =>
| .ok job log =>
let (log, jobLog) := log.split iniPos
let job := if jobLog.isEmpty then job else job.mapResult (sync := true)
(·.modifyState (.modifyLog (jobLog ++ ·)))
return (.ok job log, store)
| (.error _ log, store) =>
return .ok job log
| .error _ log =>
let (log, jobLog) := log.split iniPos
return (.ok (.error jobLog) log, store)
return .ok (.error jobLog) log

/--
Registers the job for the top-level build monitor,
Expand Down
16 changes: 16 additions & 0 deletions src/lake/Lake/Util/StoreInsts.lean
Original file line number Diff line number Diff line change
Expand Up @@ -15,19 +15,35 @@ instance [Monad m] [EqOfCmpWrt κ β cmp] : MonadDStore κ β (StateT (DRBMap κ
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

instance [MonadLiftT (ST ω) m] [Monad m] [EqOfCmpWrt κ β cmp] : MonadDStore κ β (StateRefT' ω (DRBMap κ β cmp) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

instance [Monad m] : MonadStore κ α (StateT (RBMap κ α cmp) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

instance [MonadLiftT (ST ω) m] [Monad m] : MonadStore κ α (StateRefT' ω (RBMap κ α cmp) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

instance [Monad m] : MonadStore κ α (StateT (RBArray κ α cmp) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

instance [MonadLiftT (ST ω) m] [Monad m] : MonadStore κ α (StateRefT' ω (RBArray κ α cmp) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

-- uses the eagerly specialized `RBMap` functions in `NameMap`
instance [Monad m] : MonadStore Name α (StateT (NameMap α) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

instance [MonadLiftT (ST ω) m] [Monad m] : MonadStore Name α (StateRefT' ω (NameMap α) m) where
fetch? k := return (← get).find? k
store k a := modify (·.insert k a)

@[inline] instance [MonadDStore κ β m] [t : FamilyOut β k α] : MonadStore1Of k α m where
fetch? := cast (by rw [t.family_key_eq_type]) <| fetch? (m := m) k
store a := store k <| cast t.family_key_eq_type.symm a

0 comments on commit 79d8c56

Please sign in to comment.