Library transfer.paper_model
Connecting the General Result to the Model in the Paper
From mathcomp Require Import all_ssreflect.
Require Export prosa.analysis.definitions.finish_time.
Require Export transfer_schedulability.
Elements of the EMSOFT System Model
- job precedence and delay constraints;
- system evolutions;
- a notion of schedulers.
Precedence and Delay
Job Parameters
The delay that must pass between when a job j is first scheduled and the
completion of a predecessor j' is given by job_delay j j'.
Definition of Job Readiness
Consider any kind of jobs...
... and any kind of processor state.
Suppose jobs have an arrival time, a cost, a set of predecessors, and
predecessor-specific release delays.
For a given predecessor j_pred of a job j, we say the delay has passed
at time t if j_pred completed at least job_delay j j_pred time units
earlier.
Let delay_has_passed sched j t j_pred :=
let
delta := job_delay j j_pred
in
(t ≥ delta) && completed_by sched j_pred (t - delta).
We say that a job j is released at a time t after its arrival if the
delay has passed w.r.t. each predecessor.
Let is_released sched j t :=
(job_arrival j ≤ t)
&& all (delay_has_passed sched j t) (job_predecessors j).
A job j is hence ready to be scheduled at time t if and only if it
is released and not yet complete.
#[local,program] Instance delayed_precedence_ready_instance : JobReady Job PState :=
{
job_ready sched j t := is_released sched j t && ~~ completed_by sched j t
}.
Next Obligation.
move⇒ sched j t /andP[/andP[ARR _] UNFINISHED].
by apply /andP; split.
Qed.
End DelayedPrecedenceReadiness.
System Evolutions
... the delays are given by evo_delays omega.
Consider any type of system evolutions Omega ...
... any kind of jobs, and ...
... and any kind of processor model.
A scheduler is a function that, given each job's arrival time,
predecessors, and an evolution omega, yields a schedule of the jobs.
Definition Scheduler
`{JobArrival Job} `{JobPredecessors} `{SystemEvolutions Omega}
:= Omega → schedule PState.
End SchedulerDef.
Instantiation of the EMSOFT System Model
The Set of Jobs
... which arrive according to the arrival sequence arr_seq.
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_arrivals : valid_arrival_sequence arr_seq.
For clarity, we introduce the following abbreviations:
Let job_ready (omega : Omega) : JobReady Job PState :=
@delayed_precedence_ready_instance _ _ _ (job_cost omega) _ (job_delay omega).
Worst-Case Reference Evolution
Let omega_0 denote the system evolution with maximal parameters, that
is, ...
... the job costs in omega_0 upper-bound the costs in any other
evolution, and ...
... and the delays in omega_0 similarly upper-bound the delays in any
other evolution.
As discussed in the paper, we restrict our attention to schedulers that
always produce well-formed schedules.
Hypothesis H_well_formed_A :
@valid_schedule _ _ (algA omega_0) (job_cost omega_0) _ (job_ready omega_0) arr_seq.
Hypothesis H_well_formed_B :
∀ omega, @valid_schedule _ _ (algB omega) (job_cost omega) _ (job_ready omega) arr_seq.
For brevity, recall the notion of schedulability_transferred specialized
to algA omega_0 as the reference schedule and algB omega as the online
schedule.
Definition schedulability_transferred_AB omega :=
schedulability_transferred
(algA omega_0) (algB omega) (job_cost omega_0) (job_cost omega).
Clairvoyant Results
The criterion is assumed to hold for all evolutions omega ...
... w.r.t. the respective job costs ...
... for the given, fixed job arrival times ...
with regard to online job costs. This last parameter is the
"clairvoyant" part.
The above clairvoyant criterion ensures that schedulability is transferred
in every evolution.
Theorem clairvoyant_sufficiency :
clairvoyant_criterion →
∀ omega, schedulability_transferred_AB omega.
Proof.
by move⇒ ? ?; apply: online_transfer_schedulability_criterion_sufficiency.
Qed.
Conversely, the fact that schedulability is transferred in every evolution
implies that the above criterion holds.
Theorem clairvoyant_necessity :
(∀ omega, schedulability_transferred_AB omega) →
clairvoyant_criterion.
Proof.
rewrite /schedulability_transferred_AB ⇒ TRANS omega.
by apply: online_transfer_schedulability_criterion_necessity.
Qed.
Nonclairvoyant Result
Definition nonclairvoyant_criterion :=
∀ omega,
transfer_schedulability_criterion
(algA omega_0) (algB omega)
(job_cost omega_0) (job_cost omega)
arr_seq
(job_cost omega_0).
Assuming the above criterion holds, schedulability is transferred in all
evolutions.
Theorem nonclairvoyant_sufficiency :
nonclairvoyant_criterion →
∀ omega, schedulability_transferred_AB omega.
Proof.
by move⇒ CRIT omega; apply: ref_transfer_schedulability_criterion_sufficiency.
Qed.
Results w.r.t. Finish Times
Non-Starvation Assumption
Hypothesis H_non_starvation :
∀ j,
arrives_in arr_seq j →
{ R : duration |
@job_response_time_bound _ _ (algA omega_0) (job_cost omega_0) _ j R }.
Definition of Finish Times
In the context of a given evolution omega ...
... exploit the non-starvation assumption ...
Let R := proj1_sig (H_non_starvation jf H_arrives).
Let ref_response_time_bound := proj2_sig (H_non_starvation jf H_arrives).
To define the finish time in the online schedule, we need to put in a
little more legwork, since the non-starvation assumption above applies
only to the reference schedule, but not to the online schedule. To this
end, assume that schedulability is transferred from algA omega_0 to
algB omega.
Under this assumption, R is of course also a response-time bound
w.r.t. the online schedule.
Lemma online_response_time_bound :
@job_response_time_bound _ _ (algB omega) (job_cost omega) _ jf R.
Proof. by apply/H_trans/ref_response_time_bound. Qed.
Having established that the online schedule is starvation-free, too, we
can recall the finish time in algB omega, too.
With the finish times in place for both schedules, we can state the
property that we're after: the finish time in the online schedule is no
later than in the reference schedule.
Clairvoyant Sufficiency Result for Finish Times
Under the assumption that the schedulability transfer criterion holds
w.r.t. online job costs ...
... the finish time of any job (that is part of the arrival sequence) in
any online schedule is no later than in the reference schedule.
Theorem clairvoyant_sufficiency' :
∀ omega j (IN : arrives_in arr_seq j),
online_finish_time_bounded omega j IN
(clairvoyant_sufficiency H_criterion omega).
Proof.
move⇒ omega j IN.
by apply/earliest_finish_time/clairvoyant_sufficiency/finished_at_finish_time.
Qed.
End SufficientFinishTimesClairvoyant.
Before moving on to the necessity argument for the clairvoyant case, which
will require further assumptions and legwork, we first quickly dispatch
the sufficiency argument in the case of nonclairvoyant schedulers.
The sufficiency claim for nonclairvoyant schedulers follows virtually
identically to the above theorem since the underlying argument is the
same.
Nonclairvoyant Sufficiency Result for Finish Times
Under the assumption that the schedulability transfer criterion holds
w.r.t. reference job costs ...
... the finish time of any job (that is part of the arrival sequence) in
any online schedule is no later than in the reference schedule.
Theorem nonclairvoyant_sufficiency' :
∀ omega j (IN : arrives_in arr_seq j),
online_finish_time_bounded omega j IN
(nonclairvoyant_sufficiency H_criterion omega).
Proof.
move⇒ omega j IN.
by apply/earliest_finish_time/nonclairvoyant_sufficiency/finished_at_finish_time.
Qed.
End SufficientFinishTimesNonclairvoyant.
Clairvoyant Necessity Result for Finish Times
Hypothesis H_non_starvation' :
∀ j,
arrives_in arr_seq j →
∀ omega,
{ R | @job_response_time_bound _ _ (algB omega) (job_cost omega) _ j R }.
Online Finish Times, Revisited
In the context of a given evolution omega ...
... exploit the non-starvation assumption ...
Using this notion of j's online finish time, we can re-state that it
finishes no later in the online schedule than in the offline schedule.
Definition online_finish_time_bounded' :=
online_finish_time' ≤ ref_finish_time jf H_arrives.
End FinishTimes.
Clairvoyant Necessity Theorem
Under the assumption that, for any job that arrives in the arrival
sequence, the finish time in any online schedule is no later than in the
reference schedule ...
Hypothesis H_bounded :
∀ omega j in_arrival_sequence,
online_finish_time_bounded' omega j in_arrival_sequence.
... we can show that that transfer schedulability criterion
w.r.t. online job costs holds.
Theorem clairvoyant_necessity' : clairvoyant_criterion.
Proof.
apply: clairvoyant_necessity ⇒ omega.
move⇒ j t COMP.
have [ZERO|POS] := (posnP (job_cost omega_0 j)).
{ move: (H_max_cost omega j); rewrite ZERO leqn0 ⇒ /eqP ZERO'.
by move: ZERO'; rewrite /job_cost /completed_by /job.job_cost ⇒ →. }
{ have [t' [LT SCHED']] : ∃ t' : nat, t' < t ∧ scheduled_at (algA omega_0) j t'.
{ apply: positive_service_implies_scheduled_before.
by move: POS COMP; rewrite /completed_by/service/job.job_cost/job_cost; lia. }
have IN : arrives_in arr_seq j by move: (H_well_formed_A) ⇒ [+ _]; apply.
set RTB := proj2_sig (H_non_starvation' j IN omega).
apply: completion_monotonic; last by apply/finished_at_finish_time/RTB.
move: (H_bounded omega j IN); rewrite /online_finish_time_bounded'/online_finish_time' ⇒ BOUND.
apply: (leq_trans BOUND).
exact/earliest_finish_time/COMP. }
Qed.
End NecessaryFinishTimesClairvoyant.
End EMSOFTModel.