Build Runner #
This module defines the top-level functions used to execute a Lake build, monitor its progress, and await the result.
Create a fresh build context from a workspace and a build configuration.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Unicode icons that make up the spinner in animation order.
Equations
- Lake.Monitor.spinnerFrames = #['⣾', '⣷', '⣯', '⣟', '⡿', '⢿', '⣻', '⣽']
Instances For
Context of the Lake build monitor.
- totalJobs : Nat
- out : IO.FS.Stream
- outLv : Lake.LogLevel
- failLv : Lake.LogLevel
- minAction : Lake.JobAction
- useAnsi : Bool
- showProgress : Bool
- updateFrequency : NatHow often to poll jobs (in milliseconds). 
Instances For
State of the Lake build monitor.
- jobNo : Nat
- resetCtrl : String
- lastUpdate : Nat
- spinnerIdx : Fin Lake.Monitor.spinnerFrames.size
Instances For
Monad of the Lake build monitor.
Instances For
Equations
- Lake.MonitorM.run ctx s self = self ctx s
Instances For
The ANSI escape sequence for clearing the current line and resetting the cursor back to the start.
Equations
- Lake.Ansi.resetLine = "\x1b[2K\x0d"
Instances For
Like IO.FS.Stream.flush, but ignores errors.
Equations
- Lake.flush out = EIO.catchExceptions out.flush fun (x : IO.Error) => pure ()
Instances For
Like IO.FS.Stream.putStr, but panics on errors.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lake.Monitor.print s = do let __do_lift ← read liftM (Lake.print! __do_lift.out s)
Instances For
Equations
- Lake.Monitor.flush = do let __do_lift ← read liftM (Lake.flush __do_lift.out)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
The job monitor function. An auxiliary definition for runFetchM.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Run a build function in the Workspace's context using the provided configuration.
Reports incremental build progress and build logs. In quiet mode, only reports
failing build jobs (e.g., when using -q or non-verbose --no-build).
Equations
- One or more equations did not get rendered due to their size.
Instances For
Run a build function in the Workspace's context and await the result.
Equations
- ws.runBuild build cfg = do let job ← ws.runFetchM build cfg let __discr ← liftM job.wait? match __discr with | some a => pure a | x => Lake.error "build failed"
Instances For
Produce a build job in the Lake monad's workspace and await the result.
Equations
- Lake.runBuild build cfg = do let __do_lift ← Lake.getWorkspace liftM (__do_lift.runBuild build cfg)