Library probsa.rt.model.scheduler
(* --------------------------------- Prosa ---------------------------------- *)
From prosa.model Require Import processor.ideal.
(* -------------------------------- ProBsa --------------------------------- *)
From probsa.rt Require Import behavior.response_time.
(* ------------------------------ Definitions ------------------------------- *)
From prosa.model Require Import processor.ideal.
(* -------------------------------- ProBsa --------------------------------- *)
From probsa.rt Require Import behavior.response_time.
(* ------------------------------ Definitions ------------------------------- *)
Consider any type of job.
We define a scheduler with hardcoded arrivals and costs as a
function that receives two functions -- one describing job
arrivals π : Job β option instant and another describing job
costs π : Job β option work -- and returns a schedule (which
is a function that maps a time instant to a processor state at
this time instant). Note that the functions do not depend on Ο
and, hence, are fixed. The motivation behind such a scheduler is to encode a
scheduling algorithm that makes its decisions based only on the
information provided by inputs from π and π.
Definition schedulerππ : Type :=
((* π *) Job β option instant) β ((* π *) Job β option work)
β instant β ideal.processor_state Job.
((* π *) Job β option instant) β ((* π *) Job β option work)
β instant β ideal.processor_state Job.
The following definitions just mirror the existing definitions.
The only difference is that the proposed definitions use the
scheduler instead of a generic schedule.
Definition serviceππ : Type :=
((* π *) Job β option instant) β ((* π *) Job β option work)
β Job β instant β work.
Definition completed_byππ : Type :=
((* π *) Job β option instant) β ((* π *) Job β option work)
β Job β instant β bool.
Definition response_timeππ : Type :=
((* π *) Job β option instant) β ((* π *) Job β option work)
β Job β etime.
Definition schedulerππ_to_serviceππ (ΞΆ : schedulerππ) : serviceππ :=
fun π π j t β service (ΞΆ π π) j t.
Definition schedulerππ_to_completedππ (ΞΆ : schedulerππ) : completed_byππ :=
fun π π j t β π j β¨<=β© service (ΞΆ π π) j t.
Given a scheduler, we define a function that maps π : Job β
option instant and π : Job β option work to a function that
can compute the response time of any given job.
Section ResponseTimeFromScheduler.
Variable ΞΆ : schedulerππ.
Let completed_by := schedulerππ_to_completedππ ΞΆ.
Definition completedππ_is_dec :
β π π (j : Job) (compl_time : instant),
completed_by π π j compl_time β¨ Β¬ completed_by π π j compl_time.
Definition min_completion_time π π (j : Job) :=
LPO_min (fun compl_time β completed_by π π j compl_time) (completedππ_is_dec π π j).
Definition schedulerππ_to_rtππ : response_timeππ :=
fun π π j β
match π j with
| None β Undef
| Some ar β
match min_completion_time π π j, horizon with
| inright _, _ β Infty
| inleft (exist compl_time _), None β Fin (compl_time - ar)
| inleft (exist compl_time _), Some h β
if (compl_time β₯ h)%nat then Infty else Fin (compl_time - ar)
end
end.
End ResponseTimeFromScheduler.
End SchedulerHardcoded.
Variable ΞΆ : schedulerππ.
Let completed_by := schedulerππ_to_completedππ ΞΆ.
Definition completedππ_is_dec :
β π π (j : Job) (compl_time : instant),
completed_by π π j compl_time β¨ Β¬ completed_by π π j compl_time.
Definition min_completion_time π π (j : Job) :=
LPO_min (fun compl_time β completed_by π π j compl_time) (completedππ_is_dec π π j).
Definition schedulerππ_to_rtππ : response_timeππ :=
fun π π j β
match π j with
| None β Undef
| Some ar β
match min_completion_time π π j, horizon with
| inright _, _ β Infty
| inleft (exist compl_time _), None β Fin (compl_time - ar)
| inleft (exist compl_time _), Some h β
if (compl_time β₯ h)%nat then Infty else Fin (compl_time - ar)
end
end.
End ResponseTimeFromScheduler.
End SchedulerHardcoded.
Section Schedule.
Context {Ξ©} {ΞΌ : measure Ξ©}.
Context {Job : JobType}
{job_arrival : JobArrivalRV Job Ξ© ΞΌ}
{job_cost : JobCostRV Job Ξ© ΞΌ}.
Context {Ξ©} {ΞΌ : measure Ξ©}.
Context {Job : JobType}
{job_arrival : JobArrivalRV Job Ξ© ΞΌ}
{job_cost : JobCostRV Job Ξ© ΞΌ}.
If we are given a scheduler sched, then a
probabilistic schedule is defined as follows:
Ξ» (Ο : Ξ©) (t : instant),
sched (compute_arrivals Ο) (compute_costs Ο) t.
Here, compute_arrivals (compute_costs) is a function that
maps Ο β Ξ© to a function with fixed arrivals (costs).
Similarly to arival_sequence, pr_schedule is an arrow type
(instant β PState), Coq cannot automatically derive the right
type. Therefore, we need to steer Coq's coercion and type
systems towards the right type via the annotation B :=
Equality.clone _ _.
Definition compute_pr_schedule (sched : @schedulerππ Job) : pr_schedule ΞΌ (ideal.processor_state Job) :=
mkRvar _ (B := Equality.clone _ _)
(fun Ο t β
sched (compute_arrivals Ο) (compute_costs Ο) t
).
End Schedule.
mkRvar _ (B := Equality.clone _ _)
(fun Ο t β
sched (compute_arrivals Ο) (compute_costs Ο) t
).
End Schedule.
Section RTMonotonicScheduler.
Context {Ξ©} {ΞΌ : measure Ξ©}.
Context {Job : JobType}
{job_arrival : JobArrivalRV Job Ξ© ΞΌ}
{job_cost : JobCostRV Job Ξ© ΞΌ}.
Variable horizon : option instant.
Context {Ξ©} {ΞΌ : measure Ξ©}.
Context {Job : JobType}
{job_arrival : JobArrivalRV Job Ξ© ΞΌ}
{job_cost : JobCostRV Job Ξ© ΞΌ}.
Variable horizon : option instant.
Consider a scheduling algorithm ΞΆ.
For convenience, let π‘ denote an algorithm that maps
job arrivals, job costs, and a job to its response time.
We prove that the response time of a job computed in sched is
equal to the response time of the same job computed via π‘ if
run with job arrivals equal to compute_arrivals Ο and job
costs equal to compute_costs Ο.
Lemma valid_π‘ :
β (j : Job) (Ο : Ξ©),
response_time sched horizon j Ο = π‘ (compute_arrivals Ο) (compute_costs Ο) j.
β (j : Job) (Ο : Ξ©),
response_time sched horizon j Ο = π‘ (compute_arrivals Ο) (compute_costs Ο) j.
To define RT-monotonicity, we define a function that updates a
given vector with a new cost.
Definition update (π : Job β option work) (jo : Job) (c : option work) : (Job β option work) :=
fun j β if j == jo then c else π j.
fun j β if j == jo then c else π j.
A scheduler ΞΆ is response-time monotonic if, given vectors of
fixed arrival times π and job costs π, an update of the cost
of any job j_update from c1 to c2 : c1 β¨β€β© c2 cannot cause
a decrease in response time of any job j.
Definition rt_monotonic_scheduler :=
β π π (j_update j : Job), β (r : instant) (c1 c2 : option nat),
is_true (c1 β¨<=β© c2) β
exceeds (π‘ π (update π j_update c1) j) r β
exceeds (π‘ π (update π j_update c2) j) r.
End RTMonotonicScheduler.
Opaque schedulerππ_to_rtππ.
β π π (j_update j : Job), β (r : instant) (c1 c2 : option nat),
is_true (c1 β¨<=β© c2) β
exceeds (π‘ π (update π j_update c1) j) r β
exceeds (π‘ π (update π j_update c2) j) r.
End RTMonotonicScheduler.
Opaque schedulerππ_to_rtππ.