Library prosa.implementation.facts.ideal_uni.preemption_aware
Require Export prosa.implementation.facts.generic_schedule.
Require Export prosa.implementation.definitions.ideal_uni_scheduler.
Require Export prosa.analysis.facts.model.ideal.schedule.
Require Export prosa.model.schedule.limited_preemptive.
Require Export prosa.implementation.definitions.ideal_uni_scheduler.
Require Export prosa.analysis.facts.model.ideal.schedule.
Require Export prosa.model.schedule.limited_preemptive.
Properties of the Preemption-Aware Ideal Uniprocessor Scheduler
Require Import prosa.model.processor.ideal.
Require Export prosa.util.tactics.
Section NPUniprocessorScheduler.
Require Export prosa.util.tactics.
Section NPUniprocessorScheduler.
Consider any type of jobs with costs and arrival times, ...
... in the context of an ideal uniprocessor model.
Suppose we are given a valid arrival sequence of such jobs, ...
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_arrivals : valid_arrival_sequence arr_seq.
Hypothesis H_valid_arrivals : valid_arrival_sequence arr_seq.
... a non-clairvoyant readiness model, ...
Context {RM: JobReady Job (ideal.processor_state Job)}.
Hypothesis H_nonclairvoyant_job_readiness: nonclairvoyant_readiness RM.
Hypothesis H_nonclairvoyant_job_readiness: nonclairvoyant_readiness RM.
... and a preemption model.
For any given job selection policy ...
... consider the schedule produced by the preemption-aware scheduler for
the policy induced by choose_job.
Let schedule := pmc_uni_schedule arr_seq choose_job.
Let policy := allocation_at arr_seq choose_job.
Let policy := allocation_at arr_seq choose_job.
To begin with, we establish that the preemption-aware scheduler does not
induce non-work-conserving behavior.
If choose_job does not voluntarily idle the processor, ...
... then we can establish work-conservation.
First, we observe that allocation_at yields idle_state only if there are
no backlogged jobs.
Lemma allocation_at_idle:
∀ sched t,
allocation_at arr_seq choose_job sched t = idle_state →
jobs_backlogged_at arr_seq sched t = [::].
Proof.
move⇒ sched t.
elim: t ⇒ [|t _]; first by apply H_non_idling.
rewrite /allocation_at /prev_job_nonpreemptive.
elim: (sched t) ⇒ [j'|]; last by apply H_non_idling.
move⇒ SCHED.
destruct (job_ready sched j' t.+1 && ~~ job_preemptable j' (service sched j' t.+1)) ⇒ //.
by apply (H_non_idling t.+1).
Qed.
∀ sched t,
allocation_at arr_seq choose_job sched t = idle_state →
jobs_backlogged_at arr_seq sched t = [::].
Proof.
move⇒ sched t.
elim: t ⇒ [|t _]; first by apply H_non_idling.
rewrite /allocation_at /prev_job_nonpreemptive.
elim: (sched t) ⇒ [j'|]; last by apply H_non_idling.
move⇒ SCHED.
destruct (job_ready sched j' t.+1 && ~~ job_preemptable j' (service sched j' t.+1)) ⇒ //.
by apply (H_non_idling t.+1).
Qed.
As a stepping stone, we observe that the generated schedule is idle at
a time t only if there are no backlogged jobs.
Lemma idle_schedule_no_backlogged_jobs:
∀ t,
ideal_is_idle schedule t →
jobs_backlogged_at arr_seq schedule t = [::].
Proof.
move⇒ t.
rewrite /ideal_is_idle /schedule /pmc_uni_schedule /generic_schedule ⇒ /eqP.
move⇒ NONE. move: (NONE).
rewrite schedule_up_to_def ⇒ IDLE.
apply allocation_at_idle in IDLE.
rewrite -IDLE.
apply backlogged_jobs_prefix_invariance with (h := t.+1) ⇒ //.
rewrite /identical_prefix ⇒ x.
rewrite ltnS leq_eqVlt ⇒ /orP [/eqP ->|LT]; last first.
{ elim: t LT IDLE NONE ⇒ // ⇒ h IH LT_x IDLE NONE.
by apply schedule_up_to_prefix_inclusion. }
{ elim: t IDLE NONE ⇒ [IDLE _| t' _ _ ->]; last by rewrite schedule_up_to_empty.
rewrite /schedule_up_to replace_at_def.
by rewrite /allocation_at /prev_job_nonpreemptive IDLE H_non_idling. }
Qed.
∀ t,
ideal_is_idle schedule t →
jobs_backlogged_at arr_seq schedule t = [::].
Proof.
move⇒ t.
rewrite /ideal_is_idle /schedule /pmc_uni_schedule /generic_schedule ⇒ /eqP.
move⇒ NONE. move: (NONE).
rewrite schedule_up_to_def ⇒ IDLE.
apply allocation_at_idle in IDLE.
rewrite -IDLE.
apply backlogged_jobs_prefix_invariance with (h := t.+1) ⇒ //.
rewrite /identical_prefix ⇒ x.
rewrite ltnS leq_eqVlt ⇒ /orP [/eqP ->|LT]; last first.
{ elim: t LT IDLE NONE ⇒ // ⇒ h IH LT_x IDLE NONE.
by apply schedule_up_to_prefix_inclusion. }
{ elim: t IDLE NONE ⇒ [IDLE _| t' _ _ ->]; last by rewrite schedule_up_to_empty.
rewrite /schedule_up_to replace_at_def.
by rewrite /allocation_at /prev_job_nonpreemptive IDLE H_non_idling. }
Qed.
From the preceding fact, we conclude that the generated schedule is
indeed work-conserving.
Theorem np_schedule_work_conserving:
work_conserving arr_seq schedule.
Proof.
move⇒ j t ARRIVES BACKLOGGED.
move: (@ideal_proc_model_sched_case_analysis Job schedule t) ⇒ [IDLE|SCHED]; last by exact.
exfalso.
have NON_EMPTY: j \in jobs_backlogged_at arr_seq schedule t by exact: mem_backlogged_jobs.
move: (idle_schedule_no_backlogged_jobs t IDLE) ⇒ EMPTY.
by rewrite EMPTY in NON_EMPTY.
Qed.
End WorkConservation.
work_conserving arr_seq schedule.
Proof.
move⇒ j t ARRIVES BACKLOGGED.
move: (@ideal_proc_model_sched_case_analysis Job schedule t) ⇒ [IDLE|SCHED]; last by exact.
exfalso.
have NON_EMPTY: j \in jobs_backlogged_at arr_seq schedule t by exact: mem_backlogged_jobs.
move: (idle_schedule_no_backlogged_jobs t IDLE) ⇒ EMPTY.
by rewrite EMPTY in NON_EMPTY.
Qed.
End WorkConservation.
The next result establishes that the generated preemption-model-aware
schedule is structurally valid, meaning all jobs stem from the arrival
sequence and only ready jobs are scheduled.
First, any reasonable job selection policy will not create jobs "out
of thin air," i.e., if a job is selected, it was among those given
to choose from.
Second, for the schedule to be valid, we require the notion of readiness
to be consistent with the preemption model: a non-preemptive job remains
ready until (at least) the end of its current non-preemptive section.
Finally, we assume the readiness model to be non-clairvoyant.
For notational convenience, recall the definition of a prefix of the
schedule based on which the next decision is made.
Let sched_prefix t :=
if t is t'.+1 then schedule_up_to policy None t' else empty_schedule idle_state.
if t is t'.+1 then schedule_up_to policy None t' else empty_schedule idle_state.
We begin by showing that any job in the schedule must come from the arrival
sequence used to generate it.
Lemma np_schedule_jobs_from_arrival_sequence:
jobs_come_from_arrival_sequence schedule arr_seq.
Proof.
move⇒ j t; rewrite scheduled_at_def /schedule /pmc_uni_schedule /generic_schedule.
elim: t ⇒ [/eqP | t' IH /eqP].
- rewrite schedule_up_to_def /allocation_at /prev_job_nonpreemptive ⇒ IN.
move: (H_chooses_from_set _ _ _ IN).
by apply backlogged_job_arrives_in.
- rewrite schedule_up_to_def /allocation_at.
case: (prev_job_nonpreemptive (schedule_up_to _ _ t') t'.+1) ⇒ [|IN].
+ by rewrite -pred_Sn ⇒ SCHED; apply IH; apply /eqP.
+ move: (H_chooses_from_set _ _ _ IN).
by apply backlogged_job_arrives_in.
Qed.
jobs_come_from_arrival_sequence schedule arr_seq.
Proof.
move⇒ j t; rewrite scheduled_at_def /schedule /pmc_uni_schedule /generic_schedule.
elim: t ⇒ [/eqP | t' IH /eqP].
- rewrite schedule_up_to_def /allocation_at /prev_job_nonpreemptive ⇒ IN.
move: (H_chooses_from_set _ _ _ IN).
by apply backlogged_job_arrives_in.
- rewrite schedule_up_to_def /allocation_at.
case: (prev_job_nonpreemptive (schedule_up_to _ _ t') t'.+1) ⇒ [|IN].
+ by rewrite -pred_Sn ⇒ SCHED; apply IH; apply /eqP.
+ move: (H_chooses_from_set _ _ _ IN).
by apply backlogged_job_arrives_in.
Qed.
Next, we show that any job selected by the job selection policy must
be ready.
Theorem chosen_job_is_ready:
∀ j t,
choose_job t (jobs_backlogged_at arr_seq (sched_prefix t) t) == Some j →
job_ready schedule j t.
Proof.
move⇒ j t /eqP SCHED.
apply H_chooses_from_set in SCHED.
move: SCHED; rewrite mem_filter ⇒ /andP [/andP[READY _] IN].
rewrite (H_nonclairvoyant_readiness _ (sched_prefix t) j t) //.
move⇒ t' LEQt'.
rewrite /schedule /pmc_uni_schedule /generic_schedule
schedule_up_to_def /sched_prefix.
destruct t ⇒ //=.
rewrite -schedule_up_to_def.
exact: (schedule_up_to_prefix_inclusion policy).
Qed.
∀ j t,
choose_job t (jobs_backlogged_at arr_seq (sched_prefix t) t) == Some j →
job_ready schedule j t.
Proof.
move⇒ j t /eqP SCHED.
apply H_chooses_from_set in SCHED.
move: SCHED; rewrite mem_filter ⇒ /andP [/andP[READY _] IN].
rewrite (H_nonclairvoyant_readiness _ (sched_prefix t) j t) //.
move⇒ t' LEQt'.
rewrite /schedule /pmc_uni_schedule /generic_schedule
schedule_up_to_def /sched_prefix.
destruct t ⇒ //=.
rewrite -schedule_up_to_def.
exact: (schedule_up_to_prefix_inclusion policy).
Qed.
Starting from the previous result we show that, at any instant, only
a ready job can be scheduled.
Theorem jobs_must_be_ready:
jobs_must_be_ready_to_execute schedule.
Proof.
move⇒ j t.
rewrite scheduled_at_def /schedule/pmc_uni_schedule/allocation_at
/generic_schedule schedule_up_to_def.
case PREV: prev_job_nonpreemptive;
last exact: chosen_job_is_ready.
case: t PREV ⇒ // t.
rewrite /prev_job_nonpreemptive; case: (schedule_up_to) ⇒ // {}j /andP [READY _] /eqP[]<-.
erewrite (H_nonclairvoyant_readiness _ _ _ t.+1) ⇒ // t' LT.
exact: schedule_up_to_prefix_inclusion.
Qed.
jobs_must_be_ready_to_execute schedule.
Proof.
move⇒ j t.
rewrite scheduled_at_def /schedule/pmc_uni_schedule/allocation_at
/generic_schedule schedule_up_to_def.
case PREV: prev_job_nonpreemptive;
last exact: chosen_job_is_ready.
case: t PREV ⇒ // t.
rewrite /prev_job_nonpreemptive; case: (schedule_up_to) ⇒ // {}j /andP [READY _] /eqP[]<-.
erewrite (H_nonclairvoyant_readiness _ _ _ t.+1) ⇒ // t' LT.
exact: schedule_up_to_prefix_inclusion.
Qed.
Finally, we show that the generated schedule is valid.
Theorem np_schedule_valid:
valid_schedule schedule arr_seq.
Proof.
split.
- exact: np_schedule_jobs_from_arrival_sequence.
- exact: jobs_must_be_ready.
Qed.
End Validity.
valid_schedule schedule arr_seq.
Proof.
split.
- exact: np_schedule_jobs_from_arrival_sequence.
- exact: jobs_must_be_ready.
Qed.
End Validity.
Next, we observe that the resulting schedule is consistent with the
definition of "preemption times".
Again, we require that the job-selection policy is reasonable.
For notational convenience, recall the definition of a prefix of the
schedule based on which the next decision is made.
Let prefix t := if t is t'.+1 then schedule_up_to policy idle_state t' else empty_schedule idle_state.
First, we observe that non-preemptive jobs remain scheduled as long as
they are non-preemptive.
Lemma np_job_remains_scheduled:
∀ t,
prev_job_nonpreemptive (prefix t) t →
schedule_up_to policy idle_state t t = schedule_up_to policy idle_state t t.-1.
Proof.
elim ⇒ [|t _] // NP.
rewrite schedule_up_to_def /allocation_at /policy /allocation_at.
rewrite ifT // -pred_Sn.
by rewrite schedule_up_to_widen.
Qed.
∀ t,
prev_job_nonpreemptive (prefix t) t →
schedule_up_to policy idle_state t t = schedule_up_to policy idle_state t t.-1.
Proof.
elim ⇒ [|t _] // NP.
rewrite schedule_up_to_def /allocation_at /policy /allocation_at.
rewrite ifT // -pred_Sn.
by rewrite schedule_up_to_widen.
Qed.
From this, we conclude that the predicate used to determine whether the
previously scheduled job is nonpreemptive in the computation of
np_uni_schedule is consistent with the existing notion of a
preemption_time.
Lemma np_consistent:
∀ t,
prev_job_nonpreemptive (prefix t) t →
~~ preemption_time arr_seq schedule t.
Proof.
move⇒ t.
rewrite /preemption_time scheduled_job_at_def//;
last by exact/jobs_must_arrive_to_be_ready/jobs_must_be_ready.
- elim: t ⇒ [|t _]; first by rewrite /prev_job_nonpreemptive.
rewrite /schedule /pmc_uni_schedule /generic_schedule
schedule_up_to_def /prefix /allocation_at ⇒ NP.
rewrite ifT // -pred_Sn.
move: NP; rewrite /prev_job_nonpreemptive.
elim: (schedule_up_to policy idle_state t t) ⇒ // j.
have ->: (service (fun t0 : instant ⇒ schedule_up_to policy idle_state t0 t0) j t.+1 =
service (schedule_up_to policy idle_state t) j t.+1) ⇒ //.
+ rewrite /service.
apply equal_prefix_implies_same_service_during ⇒ t' /andP [_ BOUND].
rewrite (schedule_up_to_prefix_inclusion _ _ t' t) //.
+ by move⇒ /andP [? ?].
- exact: np_schedule_jobs_from_arrival_sequence.
Qed.
End PreemptionTimes.
∀ t,
prev_job_nonpreemptive (prefix t) t →
~~ preemption_time arr_seq schedule t.
Proof.
move⇒ t.
rewrite /preemption_time scheduled_job_at_def//;
last by exact/jobs_must_arrive_to_be_ready/jobs_must_be_ready.
- elim: t ⇒ [|t _]; first by rewrite /prev_job_nonpreemptive.
rewrite /schedule /pmc_uni_schedule /generic_schedule
schedule_up_to_def /prefix /allocation_at ⇒ NP.
rewrite ifT // -pred_Sn.
move: NP; rewrite /prev_job_nonpreemptive.
elim: (schedule_up_to policy idle_state t t) ⇒ // j.
have ->: (service (fun t0 : instant ⇒ schedule_up_to policy idle_state t0 t0) j t.+1 =
service (schedule_up_to policy idle_state t) j t.+1) ⇒ //.
+ rewrite /service.
apply equal_prefix_implies_same_service_during ⇒ t' /andP [_ BOUND].
rewrite (schedule_up_to_prefix_inclusion _ _ t' t) //.
+ by move⇒ /andP [? ?].
- exact: np_schedule_jobs_from_arrival_sequence.
Qed.
End PreemptionTimes.
Finally, we establish the main feature: the generated schedule respects
the preemption-model semantics.
As a minimal validity requirement (which is a part of
valid_preemption_model), we require that any job in arr_seq must
start execution to become nonpreemptive.
Hypothesis H_valid_preemption_function :
∀ j,
arrives_in arr_seq j →
job_cannot_become_nonpreemptive_before_execution j.
∀ j,
arrives_in arr_seq j →
job_cannot_become_nonpreemptive_before_execution j.
Second, for the schedule to be valid, we require the notion of readiness
to be consistent with the preemption model: a non-preemptive job remains
ready until (at least) the end of its current non-preemptive section.
Given such a valid preemption model, we establish that the generated
schedule indeed respects the preemption model semantics.
Lemma np_respects_preemption_model :
schedule_respects_preemption_model arr_seq schedule.
Proof.
move⇒ j.
elim ⇒ [| t' IH];[by rewrite service0⇒ARR /negP ?;move:(H_valid_preemption_function j ARR)|].
move⇒ ARR NP.
have: scheduled_at schedule j t'.
{ apply contraT ⇒ NOT_SCHED.
move: (not_scheduled_implies_no_service _ _ _ NOT_SCHED) ⇒ NO_SERVICE.
rewrite -(service_last_plus_before) NO_SERVICE addn0 in NP; apply IH in NP ⇒ //.
by move /negP in NOT_SCHED. }
rewrite !scheduled_at_def /schedule/pmc_uni_schedule/generic_schedule ⇒ /eqP SCHED.
rewrite -SCHED (schedule_up_to_prefix_inclusion _ _ t' t'.+1) // np_job_remains_scheduled //.
rewrite /prev_job_nonpreemptive SCHED.
rewrite (identical_prefix_service _ schedule); last by apply schedule_up_to_identical_prefix.
apply /andP; split ⇒ //.
rewrite (H_nonclairvoyant_job_readiness _ schedule _ t'.+1) //.
exact: schedule_up_to_identical_prefix.
Qed.
End PreemptionCompliance.
End NPUniprocessorScheduler.
schedule_respects_preemption_model arr_seq schedule.
Proof.
move⇒ j.
elim ⇒ [| t' IH];[by rewrite service0⇒ARR /negP ?;move:(H_valid_preemption_function j ARR)|].
move⇒ ARR NP.
have: scheduled_at schedule j t'.
{ apply contraT ⇒ NOT_SCHED.
move: (not_scheduled_implies_no_service _ _ _ NOT_SCHED) ⇒ NO_SERVICE.
rewrite -(service_last_plus_before) NO_SERVICE addn0 in NP; apply IH in NP ⇒ //.
by move /negP in NOT_SCHED. }
rewrite !scheduled_at_def /schedule/pmc_uni_schedule/generic_schedule ⇒ /eqP SCHED.
rewrite -SCHED (schedule_up_to_prefix_inclusion _ _ t' t'.+1) // np_job_remains_scheduled //.
rewrite /prev_job_nonpreemptive SCHED.
rewrite (identical_prefix_service _ schedule); last by apply schedule_up_to_identical_prefix.
apply /andP; split ⇒ //.
rewrite (H_nonclairvoyant_job_readiness _ schedule _ t'.+1) //.
exact: schedule_up_to_identical_prefix.
Qed.
End PreemptionCompliance.
End NPUniprocessorScheduler.