mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-17 16:21:46 +03:00
21 lines
382 B
Idris
21 lines
382 B
Idris
module Data.Fuel
|
|
|
|
%default total
|
|
|
|
||| Fuel for running total operations potentially indefinitely.
|
|
public export
|
|
data Fuel = Dry | More (Lazy Fuel)
|
|
|
|
||| Provide `n` units of fuel.
|
|
export
|
|
limit : Nat -> Fuel
|
|
limit Z = Dry
|
|
limit (S n) = More (limit n)
|
|
|
|
||| Provide fuel indefinitely.
|
|
||| This function is fundamentally partial.
|
|
partial
|
|
export
|
|
forever : Fuel
|
|
forever = More forever
|