Skip to content

Commit f3acab7

Browse files
algebraic-devTwoFXhargoniX
committed
feat: add async timer primitives based on libuv (#6219)
This PR adds support for `Timer` and a event loop thread that gets requests from another threads and executes it. --------- Co-authored-by: Markus Himmel <markus@himmel-villmar.de> Co-authored-by: Henrik Böving <hargonix@gmail.com>
1 parent 439b425 commit f3acab7

File tree

9 files changed

+611
-4
lines changed

9 files changed

+611
-4
lines changed

src/Std/Internal/UV.lean

Lines changed: 72 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -42,6 +42,78 @@ opaque alive : BaseIO Bool
4242

4343
end Loop
4444

45+
private opaque TimerImpl : NonemptyType.{0}
46+
47+
/--
48+
`Timer`s are used to generate `IO.Promise`s that resolve after some time.
49+
50+
A `Timer` can be in one of 3 states:
51+
- Right after construction it's initial.
52+
- While it is ticking it's running.
53+
- If it has stopped for some reason it's finished.
54+
55+
This together with whether it was set up as `repeating` with `Timer.new` determines the behavior
56+
of all functions on `Timer`s.
57+
-/
58+
def Timer : Type := TimerImpl.type
59+
60+
instance : Nonempty Timer := TimerImpl.property
61+
62+
namespace Timer
63+
64+
/--
65+
This creates a `Timer` in the initial state and doesn't run it yet.
66+
- If `repeating` is `false` this constructs a timer that resolves once after `durationMs`
67+
milliseconds, counting from when it's run.
68+
- If `repeating` is `true` this constructs a timer that resolves after multiples of `durationMs`
69+
milliseconds, counting from when it's run. Note that this includes the 0th multiple right after
70+
starting the timer. Furthermore a repeating timer will only be freed after `Timer.stop` is called.
71+
-/
72+
@[extern "lean_uv_timer_mk"]
73+
opaque mk (timeout : UInt64) (repeating : Bool) : IO Timer
74+
75+
/--
76+
This function has different behavior depending on the state and configuration of the `Timer`:
77+
- if `repeating` is `false` and:
78+
- it is initial, run it and return a new `IO.Promise` that is set to resolve once `durationMs`
79+
milliseconds have elapsed. After this `IO.Promise` is resolved the `Timer` is finished.
80+
- it is running or finished, return the same `IO.Promise` that the first call to `next` returned.
81+
- if `repeating` is `true` and:
82+
- it is initial, run it and return a new `IO.Promise` that resolves right away
83+
(as it is the 0th multiple of `durationMs`).
84+
- it is running, check whether the last returned `IO.Promise` is already resolved:
85+
- If it is, return a new `IO.Promise` that resolves upon finishing the next cycle
86+
- If it is not, return the last `IO.Promise`
87+
This ensures that the returned `IO.Promise` resolves at the next repetition of the timer.
88+
- if it is finished, return the last `IO.Promise` created by `next`. Notably this could be one
89+
that never resolves if the timer was stopped before fulfilling the last one.
90+
-/
91+
@[extern "lean_uv_timer_next"]
92+
opaque next (timer : @& Timer) : IO (IO.Promise Unit)
93+
94+
/--
95+
This function has different behavior depending on the state and configuration of the `Timer`:
96+
- If it is initial or finished this is a no-op.
97+
- If it is running and `repeating` is `false` this will delay the resolution of the timer until
98+
`durationMs` milliseconds after the call of this function.
99+
- Delay the resolution of the next tick of the timer until `durationMs` milliseconds after the
100+
call of this function, then continue normal ticking behavior from there.
101+
-/
102+
@[extern "lean_uv_timer_reset"]
103+
opaque reset (timer : @& Timer) : IO Unit
104+
105+
/--
106+
This function has different behavior depending on the state of the `Timer`:
107+
- If it is initial or finished this is a no-op.
108+
- If it is running the execution of the timer is stopped and it is put into the finished state.
109+
Note that if the last `IO.Promise` generated by `next` is unresolved and being waited
110+
on this creates a memory leak and the waiting task is not going to be awoken anymore.
111+
-/
112+
@[extern "lean_uv_timer_stop"]
113+
opaque stop (timer : @& Timer) : IO Unit
114+
115+
end Timer
116+
45117
end UV
46118
end Internal
47119
end Std

src/runtime/CMakeLists.txt

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2,7 +2,8 @@ set(RUNTIME_OBJS debug.cpp thread.cpp mpz.cpp utf8.cpp
22
object.cpp apply.cpp exception.cpp interrupt.cpp memory.cpp
33
stackinfo.cpp compact.cpp init_module.cpp load_dynlib.cpp io.cpp hash.cpp
44
platform.cpp alloc.cpp allocprof.cpp sharecommon.cpp stack_overflow.cpp
5-
process.cpp object_ref.cpp mpn.cpp mutex.cpp libuv.cpp uv/net_addr.cpp uv/event_loop.cpp)
5+
process.cpp object_ref.cpp mpn.cpp mutex.cpp libuv.cpp uv/net_addr.cpp uv/event_loop.cpp
6+
uv/timer.cpp)
67
add_library(leanrt_initial-exec STATIC ${RUNTIME_OBJS})
78
set_target_properties(leanrt_initial-exec PROPERTIES
89
ARCHIVE_OUTPUT_DIRECTORY ${CMAKE_CURRENT_BINARY_DIR})

src/runtime/libuv.cpp

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,6 +14,7 @@ namespace lean {
1414
#include <uv.h>
1515

1616
extern "C" void initialize_libuv() {
17+
initialize_libuv_timer();
1718
initialize_libuv_loop();
1819

1920
lthread([]() { event_loop_run_loop(&global_ev); });
@@ -33,4 +34,4 @@ extern "C" LEAN_EXPORT lean_obj_res lean_libuv_version(lean_obj_arg o) {
3334
}
3435

3536
#endif
36-
}
37+
}

src/runtime/libuv.h

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,7 @@ Author: Markus Himmel, Sofia Rodrigues
77
#pragma once
88
#include <lean/lean.h>
99
#include "runtime/uv/event_loop.h"
10+
#include "runtime/uv/timer.h"
1011
#include "runtime/alloc.h"
1112
#include "runtime/io.h"
1213
#include "runtime/utf8.h"
@@ -26,4 +27,4 @@ extern "C" void initialize_libuv();
2627
// General LibUV functions.
2728
extern "C" LEAN_EXPORT lean_obj_res lean_libuv_version(lean_obj_arg);
2829

29-
}
30+
}

src/runtime/object.h

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -467,6 +467,11 @@ inline obj_res st_ref_set(b_obj_arg r, obj_arg v, obj_arg w) { return lean_st_re
467467
inline obj_res st_ref_reset(b_obj_arg r, obj_arg w) { return lean_st_ref_reset(r, w); }
468468
inline obj_res st_ref_swap(b_obj_arg r, obj_arg v, obj_arg w) { return lean_st_ref_swap(r, v, w); }
469469

470+
471+
extern "C" LEAN_EXPORT obj_res lean_io_promise_new(obj_arg);
472+
extern "C" LEAN_EXPORT obj_res lean_io_promise_resolve(obj_arg value, b_obj_arg promise, obj_arg);
473+
extern "C" LEAN_EXPORT obj_res lean_io_promise_result(obj_arg promise);
474+
470475
// =======================================
471476
// Module initialization/finalization
472477
void initialize_object();

src/runtime/uv/event_loop.cpp

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -140,4 +140,4 @@ extern "C" LEAN_EXPORT lean_obj_res lean_uv_event_loop_alive(obj_arg /* w */ ) {
140140

141141
#endif
142142

143-
}
143+
}

src/runtime/uv/timer.cpp

Lines changed: 254 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,254 @@
1+
/*
2+
Copyright (c) 2024 Lean FRO, LLC. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
5+
Author: Sofia Rodrigues, Henrik Böving
6+
*/
7+
#include "runtime/uv/timer.h"
8+
9+
namespace lean {
10+
#ifndef LEAN_EMSCRIPTEN
11+
12+
using namespace std;
13+
14+
// The finalizer of the `Timer`.
15+
void lean_uv_timer_finalizer(void* ptr) {
16+
lean_uv_timer_object * timer = (lean_uv_timer_object*) ptr;
17+
18+
if (timer->m_promise != NULL) {
19+
lean_dec(timer->m_promise);
20+
}
21+
22+
event_loop_lock(&global_ev);
23+
24+
uv_close((uv_handle_t*)timer->m_uv_timer, [](uv_handle_t* handle) {
25+
free(handle);
26+
});
27+
28+
event_loop_unlock(&global_ev);
29+
30+
free(timer);
31+
}
32+
33+
void initialize_libuv_timer() {
34+
g_uv_timer_external_class = lean_register_external_class(lean_uv_timer_finalizer, [](void* obj, lean_object* f) {
35+
if (((lean_uv_timer_object*)obj)->m_promise != NULL) {
36+
lean_inc(f);
37+
lean_apply_1(f, ((lean_uv_timer_object*)obj)->m_promise);
38+
}
39+
});
40+
}
41+
42+
void handle_timer_event(uv_timer_t* handle) {
43+
lean_object * obj = (lean_object*)handle->data;
44+
lean_uv_timer_object * timer = lean_to_uv_timer(obj);
45+
// handle_timer_event may only be called while the timer is running, this means the promise must
46+
// not be NULL.
47+
lean_assert(timer->m_state == TIMER_STATE_RUNNING);
48+
lean_assert(timer->m_promise != NULL);
49+
50+
if (timer->m_repeating) {
51+
if (lean_io_get_task_state_core(timer->m_promise) != 2) {
52+
lean_object* res = lean_io_promise_resolve(lean_box(0), timer->m_promise, lean_io_mk_world());
53+
lean_dec(res);
54+
}
55+
} else {
56+
lean_assert(lean_io_get_task_state_core(timer->m_promise) != 2);
57+
uv_timer_stop(timer->m_uv_timer);
58+
timer->m_state = TIMER_STATE_FINISHED;
59+
60+
lean_object* res = lean_io_promise_resolve(lean_box(0), timer->m_promise, lean_io_mk_world());
61+
lean_dec(res);
62+
63+
// The loop does not need to keep the timer alive anymore.
64+
lean_dec(obj);
65+
}
66+
}
67+
68+
/* Std.Internal.UV.Timer.mk (timeout : UInt64) (repeating : Bool) : IO Timer */
69+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_mk(uint64_t timeout, uint8_t repeating, obj_arg /* w */) {
70+
lean_uv_timer_object * timer = (lean_uv_timer_object*)malloc(sizeof(lean_uv_timer_object));
71+
timer->m_timeout = timeout;
72+
timer->m_repeating = repeating;
73+
timer->m_state = TIMER_STATE_INITIAL;
74+
timer->m_promise = NULL;
75+
76+
uv_timer_t * uv_timer = (uv_timer_t*)malloc(sizeof(uv_timer_t));
77+
78+
event_loop_lock(&global_ev);
79+
int result = uv_timer_init(global_ev.loop, uv_timer);
80+
event_loop_unlock(&global_ev);
81+
82+
if (result != 0) {
83+
free(uv_timer);
84+
free(timer);
85+
std::string err = std::string("failed to initialize timer: ") + uv_strerror(result);
86+
return io_result_mk_error(err.c_str());
87+
}
88+
89+
timer->m_uv_timer = uv_timer;
90+
91+
lean_object * obj = lean_uv_timer_new(timer);
92+
lean_mark_mt(obj);
93+
timer->m_uv_timer->data = obj;
94+
95+
return lean_io_result_mk_ok(obj);
96+
}
97+
98+
/* Std.Internal.UV.Timer.next (timer : @& Timer) : IO (IO.Promise Unit) */
99+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_next(b_obj_arg obj, obj_arg /* w */ ) {
100+
lean_uv_timer_object * timer = lean_to_uv_timer(obj);
101+
102+
auto create_promise = []() {
103+
lean_object * prom_res = lean_io_promise_new(lean_io_mk_world());
104+
lean_object * promise = lean_ctor_get(prom_res, 0);
105+
lean_inc(promise);
106+
lean_dec(prom_res);
107+
108+
return promise;
109+
};
110+
111+
auto setup_timer = [create_promise, obj, timer]() {
112+
lean_assert(timer->m_promise == NULL);
113+
timer->m_promise = create_promise();
114+
timer->m_state = TIMER_STATE_RUNNING;
115+
116+
// The event loop must keep the timer alive for the duration of the run time.
117+
lean_inc(obj);
118+
119+
event_loop_lock(&global_ev);
120+
121+
int result = uv_timer_start(
122+
timer->m_uv_timer,
123+
handle_timer_event,
124+
timer->m_repeating ? 0 : timer->m_timeout,
125+
timer->m_repeating ? timer->m_timeout : 0
126+
);
127+
128+
event_loop_unlock(&global_ev);
129+
130+
if (result != 0) {
131+
lean_dec(obj);
132+
std::string err = std::string("failed to initialize timer: ") + uv_strerror(result);
133+
return io_result_mk_error(err.c_str());
134+
} else {
135+
lean_inc(timer->m_promise);
136+
return lean_io_result_mk_ok(timer->m_promise);
137+
}
138+
};
139+
140+
if (timer->m_repeating) {
141+
switch (timer->m_state) {
142+
case TIMER_STATE_INITIAL:
143+
{
144+
return setup_timer();
145+
}
146+
case TIMER_STATE_RUNNING:
147+
{
148+
lean_assert(timer->m_promise != NULL);
149+
// 2 indicates finished
150+
if (lean_io_get_task_state_core(timer->m_promise) == 2) {
151+
lean_dec(timer->m_promise);
152+
timer->m_promise = create_promise();
153+
lean_inc(timer->m_promise);
154+
return lean_io_result_mk_ok(timer->m_promise);
155+
} else {
156+
lean_inc(timer->m_promise);
157+
return lean_io_result_mk_ok(timer->m_promise);
158+
}
159+
}
160+
case TIMER_STATE_FINISHED:
161+
{
162+
lean_assert(timer->m_promise != NULL);
163+
lean_inc(timer->m_promise);
164+
return lean_io_result_mk_ok(timer->m_promise);
165+
}
166+
}
167+
} else {
168+
if (timer->m_state == TIMER_STATE_INITIAL) {
169+
return setup_timer();
170+
} else {
171+
lean_assert(timer->m_promise != NULL);
172+
173+
lean_inc(timer->m_promise);
174+
return lean_io_result_mk_ok(timer->m_promise);
175+
}
176+
}
177+
}
178+
179+
/* Std.Internal.UV.Timer.reset (timer : @& Timer) : IO Unit */
180+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_reset(b_obj_arg obj, obj_arg /* w */ ) {
181+
lean_uv_timer_object * timer = lean_to_uv_timer(obj);
182+
183+
if (timer->m_state == TIMER_STATE_RUNNING) {
184+
lean_assert(timer->m_promise != NULL);
185+
186+
event_loop_lock(&global_ev);
187+
188+
uv_timer_stop(timer->m_uv_timer);
189+
190+
int result = uv_timer_start(
191+
timer->m_uv_timer,
192+
handle_timer_event,
193+
timer->m_timeout,
194+
timer->m_repeating ? timer->m_timeout : 0
195+
);
196+
197+
event_loop_unlock(&global_ev);
198+
199+
if (result != 0) {
200+
return io_result_mk_error("failed to restart uv_timer");
201+
} else {
202+
return lean_io_result_mk_ok(lean_box(0));
203+
}
204+
} else {
205+
return lean_io_result_mk_ok(lean_box(0));
206+
}
207+
}
208+
209+
/* Std.Internal.UV.Timer.stop (timer : @& Timer) : IO Unit */
210+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_stop(b_obj_arg obj, obj_arg /* w */) {
211+
lean_uv_timer_object * timer = lean_to_uv_timer(obj);
212+
213+
if (timer->m_state == TIMER_STATE_RUNNING) {
214+
lean_assert(timer->m_promise != NULL);
215+
216+
event_loop_lock(&global_ev);
217+
218+
uv_timer_stop(timer->m_uv_timer);
219+
220+
event_loop_unlock(&global_ev);
221+
222+
timer->m_state = TIMER_STATE_FINISHED;
223+
224+
// The loop does not need to keep the timer alive anymore.
225+
lean_dec(obj);
226+
227+
return lean_io_result_mk_ok(lean_box(0));
228+
} else {
229+
return lean_io_result_mk_ok(lean_box(0));
230+
}
231+
}
232+
233+
#else
234+
235+
void lean_uv_timer_finalizer(void* ptr);
236+
237+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_mk(uint64_t timeout, uint8_t repeating, obj_arg /* w */) {
238+
return io_result_mk_error("lean_uv_timer_mk is not supported");
239+
}
240+
241+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_next(b_obj_arg timer, obj_arg /* w */ ) {
242+
return io_result_mk_error("lean_uv_timer_next is not supported");
243+
}
244+
245+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_reset(b_obj_arg timer, obj_arg /* w */ ) {
246+
return io_result_mk_error("lean_uv_timer_reset is not supported");
247+
}
248+
249+
extern "C" LEAN_EXPORT lean_obj_res lean_uv_timer_stop(b_obj_arg timer, obj_arg /* w */ ) {
250+
return io_result_mk_error("lean_uv_timer_stop is not supported");
251+
}
252+
253+
#endif
254+
}

0 commit comments

Comments
 (0)