Library prosa.analysis.abstract.ideal.iw_instantiation
Require Export prosa.analysis.abstract.ideal.abstract_seq_rta.
Require Export prosa.analysis.definitions.priority_inversion.
Require Export prosa.analysis.definitions.work_bearing_readiness.
Require Export prosa.analysis.definitions.interference.
Require Export prosa.analysis.facts.busy_interval.carry_in.
Require Export prosa.analysis.abstract.iw_auxiliary.
Require Export prosa.analysis.facts.priority.classes.
Require Export prosa.analysis.definitions.priority_inversion.
Require Export prosa.analysis.definitions.work_bearing_readiness.
Require Export prosa.analysis.definitions.interference.
Require Export prosa.analysis.facts.busy_interval.carry_in.
Require Export prosa.analysis.abstract.iw_auxiliary.
Require Export prosa.analysis.facts.priority.classes.
Throughout this file, we assume ideal uni-processor schedules.
Require Import prosa.model.processor.ideal.
Require Export prosa.analysis.facts.model.ideal.priority_inversion.
Require Export prosa.analysis.facts.model.ideal.priority_inversion.
JLFP instantiation of Interference and Interfering Workload for ideal uni-processor.
In this module we instantiate functions Interference and Interfering Workload for ideal uni-processor schedulers with an arbitrary JLFP-policy that satisfies the sequential-tasks hypothesis. We also prove equivalence of Interference and Interfering Workload to the more conventional notions of service or workload.
Consider any type of tasks ...
... and any type of jobs associated with these tasks.
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Consider any valid arrival sequence with consistent arrivals...
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_arrival_sequence : valid_arrival_sequence arr_seq.
Hypothesis H_valid_arrival_sequence : valid_arrival_sequence arr_seq.
... and any ideal uni-processor schedule of this arrival
sequence...
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_jobs_come_from_arrival_sequence :
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_jobs_come_from_arrival_sequence :
jobs_come_from_arrival_sequence sched arr_seq.
... where jobs do not execute before their arrival or after
completion.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Consider a JLFP-policy that indicates a higher-or-equal priority
relation, and assume that this relation is reflexive and
transitive.
Context {JLFP : JLFP_policy Job}.
Hypothesis H_priority_is_reflexive : reflexive_job_priorities JLFP.
Hypothesis H_priority_is_transitive : transitive_job_priorities JLFP.
Hypothesis H_priority_is_reflexive : reflexive_job_priorities JLFP.
Hypothesis H_priority_is_transitive : transitive_job_priorities JLFP.
Let tsk be any task.
We also assume that the policy respects sequential tasks,
meaning that later-arrived jobs of a task don't have higher
priority than earlier-arrived jobs of the same task.
Interference and Interfering Workload
In the following, we introduce definitions of interference, interfering workload and a function that bounds cumulative interference.Instantiation of Interference
#[local,program] Instance ideal_jlfp_interference : Interference Job :=
{
interference (j : Job) (t : instant) :=
priority_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
}.
{
interference (j : Job) (t : instant) :=
priority_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
}.
Instantiation of Interfering Workload
#[local,program] Instance ideal_jlfp_interfering_workload : InterferingWorkload Job :=
{
interfering_workload (j : Job) (t : instant) :=
priority_inversion arr_seq sched j t
+ other_hep_jobs_interfering_workload arr_seq j t
}.
{
interfering_workload (j : Job) (t : instant) :=
priority_inversion arr_seq sched j t
+ other_hep_jobs_interfering_workload arr_seq j t
}.
Auxiliary definitions
Definition workload_of_another_hep_jobs (j : Job) (t1 t2 : instant) :=
workload_of_jobs (fun jhp ⇒ another_hep_job jhp j) (arrivals_between arr_seq t1 t2).
workload_of_jobs (fun jhp ⇒ another_hep_job jhp j) (arrivals_between arr_seq t1 t2).
... and service of all other jobs with higher or equal priority.
Definition service_of_another_hep_jobs (j : Job) (t1 t2 : instant) :=
service_of_jobs sched (fun jhp ⇒ another_hep_job jhp j) (arrivals_between arr_seq t1 t2) t1 t2.
service_of_jobs sched (fun jhp ⇒ another_hep_job jhp j) (arrivals_between arr_seq t1 t2) t1 t2.
Similarly, we recall notions of service of higher or equal
priority jobs from other tasks.
Definition service_of_another_task_hep_job (j : Job) (t1 t2 : instant) :=
service_of_jobs sched (fun jhp ⇒ another_task_hep_job jhp j) (arrivals_between arr_seq t1 t2) t1 t2.
service_of_jobs sched (fun jhp ⇒ another_task_hep_job jhp j) (arrivals_between arr_seq t1 t2) t1 t2.
Equivalences
In this section we prove useful equivalences between the definitions obtained by instantiation of definitions from the Abstract RTA module (interference and interfering workload) and definitions corresponding to the conventional concepts.
In the following subsection, we prove properties of the
introduced functions under the assumption that the schedule is
idle.
Consider a time instant t ...
... and assume that the schedule is idle at t.
We prove that in this case: ...
... there is no interference from higher-or-equal priority
jobs, ...
Lemma idle_implies_no_hep_job_interference :
∀ j, ¬ another_hep_job_interference arr_seq sched j t.
Proof.
move⇒ j /hasP[j' /[!mem_filter]/andP[]].
by rewrite /receives_service_at ideal_not_idle_implies_sched.
Qed.
∀ j, ¬ another_hep_job_interference arr_seq sched j t.
Proof.
move⇒ j /hasP[j' /[!mem_filter]/andP[]].
by rewrite /receives_service_at ideal_not_idle_implies_sched.
Qed.
... there is no interference from higher-or-equal priority
jobs from another task, ...
Lemma idle_implies_no_hep_task_interference :
∀ j, ¬ another_task_hep_job_interference arr_seq sched j t.
Proof.
move⇒ j /hasP[j' /[!mem_filter]/andP[]].
by rewrite /receives_service_at ideal_not_idle_implies_sched.
Qed.
∀ j, ¬ another_task_hep_job_interference arr_seq sched j t.
Proof.
move⇒ j /hasP[j' /[!mem_filter]/andP[]].
by rewrite /receives_service_at ideal_not_idle_implies_sched.
Qed.
... there is no interference, ...
Lemma idle_implies_no_interference :
∀ j, ¬ interference j t.
Proof.
move ⇒ j.
rewrite /interference /ideal_jlfp_interference ⇒ /orP [PI | HEPI].
- have [j' schj'] : ∃ j' : Job, scheduled_at sched j' t.
exact/priority_inversion_scheduled_at.
exact: ideal_sched_implies_not_idle schj' H_idle.
- exact: idle_implies_no_hep_job_interference HEPI.
Qed.
∀ j, ¬ interference j t.
Proof.
move ⇒ j.
rewrite /interference /ideal_jlfp_interference ⇒ /orP [PI | HEPI].
- have [j' schj'] : ∃ j' : Job, scheduled_at sched j' t.
exact/priority_inversion_scheduled_at.
exact: ideal_sched_implies_not_idle schj' H_idle.
- exact: idle_implies_no_hep_job_interference HEPI.
Qed.
... as well as no interference for tsk. Recall that the
additional argument upper_bound is an artificial horizon
needed needed to make task interference function
constructive. For more details, refer to the original
description of the function.
Lemma idle_implies_no_task_interference :
∀ j, ¬ task_interference arr_seq sched j t.
Proof.
move⇒ j; rewrite /task_interference /cond_interference.
have ->: interference j t = false; last by rewrite andbF.
apply/negbTE/negP/idle_implies_no_interference.
Qed.
End IdleSchedule.
∀ j, ¬ task_interference arr_seq sched j t.
Proof.
move⇒ j; rewrite /task_interference /cond_interference.
have ->: interference j t = false; last by rewrite andbF.
apply/negbTE/negP/idle_implies_no_interference.
Qed.
End IdleSchedule.
Next, we prove properties of the introduced functions under
the assumption that the scheduler is not idle.
Consider a job j of task tsk. In this subsection, job
j is deemed to be the main job with respect to which the
functions are computed.
Consider a time instant t.
First, consider a case when some jobs is scheduled at time t.
Under the stated assumptions, we show that the
interference from another higher-or-equal priority job is
equivalent to the relation another_hep_job.
Lemma interference_ahep_equiv_ahep :
another_hep_job_interference arr_seq sched j t = another_hep_job j' j.
Proof.
apply/idP/idP ⇒ [/hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP] | AHEP].
- apply service_at_implies_scheduled_at in PSERV.
by have → := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_sched PSERV.
- apply/hasP; ∃ j' ⇒ //; rewrite mem_filter; apply/andP; split.
by rewrite /receives_service_at service_at_is_scheduled_at H_sched.
apply: arrived_between_implies_in_arrivals ⇒ //.
apply/andP; split⇒ [//|].
by apply H_jobs_must_arrive_to_execute in H_sched; rewrite ltnS.
Qed.
another_hep_job_interference arr_seq sched j t = another_hep_job j' j.
Proof.
apply/idP/idP ⇒ [/hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP] | AHEP].
- apply service_at_implies_scheduled_at in PSERV.
by have → := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_sched PSERV.
- apply/hasP; ∃ j' ⇒ //; rewrite mem_filter; apply/andP; split.
by rewrite /receives_service_at service_at_is_scheduled_at H_sched.
apply: arrived_between_implies_in_arrivals ⇒ //.
apply/andP; split⇒ [//|].
by apply H_jobs_must_arrive_to_execute in H_sched; rewrite ltnS.
Qed.
Similarly, we show that the interference from another
higher-or-equal priority job from another task is
equivalent to the relation another_task_hep_job.
Lemma interference_athep_equiv_athep :
another_task_hep_job_interference arr_seq sched j t = another_task_hep_job j' j.
Proof.
apply/idP/idP ⇒ [/hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP] | AHEP].
- apply service_at_implies_scheduled_at in PSERV.
by have → := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_sched PSERV.
- apply/hasP; ∃ j' ⇒ //; rewrite mem_filter; apply/andP; split.
by rewrite /receives_service_at service_at_is_scheduled_at H_sched.
apply: arrived_between_implies_in_arrivals ⇒ //.
apply/andP; split⇒ [//|].
by apply H_jobs_must_arrive_to_execute in H_sched; rewrite ltnS.
Qed.
End SomeJobIsScheduled.
another_task_hep_job_interference arr_seq sched j t = another_task_hep_job j' j.
Proof.
apply/idP/idP ⇒ [/hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP] | AHEP].
- apply service_at_implies_scheduled_at in PSERV.
by have → := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_sched PSERV.
- apply/hasP; ∃ j' ⇒ //; rewrite mem_filter; apply/andP; split.
by rewrite /receives_service_at service_at_is_scheduled_at H_sched.
apply: arrived_between_implies_in_arrivals ⇒ //.
apply/andP; split⇒ [//|].
by apply H_jobs_must_arrive_to_execute in H_sched; rewrite ltnS.
Qed.
End SomeJobIsScheduled.
Then there is no interference from higher-or-equal
priority jobs at time t.
Lemma interference_ahep_job_eq_false :
¬ another_hep_job_interference arr_seq sched j t.
Proof.
move⇒ /hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP].
apply service_at_implies_scheduled_at in PSERV.
have EQ := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j_sched PSERV; subst jhp.
by apply another_hep_job_antireflexive in AHEP.
Qed.
End JIsScheduled.
¬ another_hep_job_interference arr_seq sched j t.
Proof.
move⇒ /hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP].
apply service_at_implies_scheduled_at in PSERV.
have EQ := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j_sched PSERV; subst jhp.
by apply another_hep_job_antireflexive in AHEP.
Qed.
End JIsScheduled.
In the next subsection, we consider a case when a job j'
from the same task (as job j) is scheduled.
Variable j' : Job.
Hypothesis H_j'_tsk : job_of_task tsk j'.
Hypothesis H_j'_sched : scheduled_at sched j' t.
Hypothesis H_j'_tsk : job_of_task tsk j'.
Hypothesis H_j'_sched : scheduled_at sched j' t.
Then we show that interference from higher-or-equal
priority jobs from another task is false.
Lemma interference_athep_eq_false :
¬ another_task_hep_job_interference arr_seq sched j t.
Proof.
move⇒ /hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP].
apply service_at_implies_scheduled_at in PSERV.
have EQ := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j'_sched PSERV; subst jhp.
by eapply another_task_hep_job_taskwise_antireflexive in AHEP.
Qed.
¬ another_task_hep_job_interference arr_seq sched j t.
Proof.
move⇒ /hasP[jhp /[!mem_filter]/andP[PSERV IN] AHEP].
apply service_at_implies_scheduled_at in PSERV.
have EQ := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j'_sched PSERV; subst jhp.
by eapply another_task_hep_job_taskwise_antireflexive in AHEP.
Qed.
Similarly, there is no task interference, since in order
to incur the task interference, a job from a distinct task
must be scheduled.
Lemma task_interference_eq_false :
¬ task_interference arr_seq sched j t.
Proof.
move ⇒ /andP [+ _]; rewrite /non_self /task_scheduled_at.
rewrite task_served_eq_task_scheduled //=; erewrite job_of_scheduled_task ⇒ //.
move: H_j_tsk H_j'_tsk; rewrite /job_of_task ⇒ /eqP → /eqP →.
by rewrite eq_refl.
Qed.
End FromSameTask.
¬ task_interference arr_seq sched j t.
Proof.
move ⇒ /andP [+ _]; rewrite /non_self /task_scheduled_at.
rewrite task_served_eq_task_scheduled //=; erewrite job_of_scheduled_task ⇒ //.
move: H_j_tsk H_j'_tsk; rewrite /job_of_task ⇒ /eqP → /eqP →.
by rewrite eq_refl.
Qed.
End FromSameTask.
In the next subsection, we consider a case when a job j'
from a task other than j's task is scheduled.
Variable j' : Job.
Hypothesis H_j'_not_tsk : ~~ job_of_task tsk j'.
Hypothesis H_j'_sched : scheduled_at sched j' t.
Hypothesis H_j'_not_tsk : ~~ job_of_task tsk j'.
Hypothesis H_j'_sched : scheduled_at sched j' t.
We prove that then j incurs higher-or-equal priority
interference from another task iff j' has
higher-or-equal priority than j.
Lemma sched_at_implies_interference_athep_eq_hep :
another_task_hep_job_interference arr_seq sched j t = hep_job j' j.
Proof.
apply/idP/idP ⇒ [/hasP[j'' /[!mem_filter]/andP[RSERV IN] AHEP] | HEP].
- apply service_at_implies_scheduled_at in RSERV.
have → := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j'_sched RSERV.
by move: AHEP ⇒ /andP[].
- apply/hasP; ∃ j'; [rewrite !mem_filter|]; apply/andP; split.
+ by rewrite /receives_service_at service_at_is_scheduled_at H_j'_sched.
+ apply: arrived_between_implies_in_arrivals ⇒ //.
apply/andP; split⇒ [//|].
by apply H_jobs_must_arrive_to_execute in H_j'_sched; rewrite ltnS.
+ by [].
+ by apply: contraNN H_j'_not_tsk ⇒ /eqP; rewrite /job_of_task ⇒ →.
Qed.
another_task_hep_job_interference arr_seq sched j t = hep_job j' j.
Proof.
apply/idP/idP ⇒ [/hasP[j'' /[!mem_filter]/andP[RSERV IN] AHEP] | HEP].
- apply service_at_implies_scheduled_at in RSERV.
have → := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j'_sched RSERV.
by move: AHEP ⇒ /andP[].
- apply/hasP; ∃ j'; [rewrite !mem_filter|]; apply/andP; split.
+ by rewrite /receives_service_at service_at_is_scheduled_at H_j'_sched.
+ apply: arrived_between_implies_in_arrivals ⇒ //.
apply/andP; split⇒ [//|].
by apply H_jobs_must_arrive_to_execute in H_j'_sched; rewrite ltnS.
+ by [].
+ by apply: contraNN H_j'_not_tsk ⇒ /eqP; rewrite /job_of_task ⇒ →.
Qed.
Hence, if we assume that j' has higher-or-equal priority, ...
... we are able to show that j incurs higher-or-equal
priority interference from another task.
Lemma sched_athep_implies_interference_athep :
another_task_hep_job_interference arr_seq sched j t.
Proof.
by rewrite sched_at_implies_interference_athep_eq_hep.
Qed.
another_task_hep_job_interference arr_seq sched j t.
Proof.
by rewrite sched_at_implies_interference_athep_eq_hep.
Qed.
Moreover, in this case, task tsk also incurs interference.
Lemma sched_athep_implies_task_interference :
task_interference arr_seq sched j t.
Proof.
apply/andP; split.
- rewrite /non_self task_served_eq_task_scheduled ⇒ //.
apply: job_of_other_task_scheduled ⇒ //.
by move: H_j_tsk ⇒ /eqP → .
- apply/orP; right.
apply/hasP; ∃ j'.
+ by apply scheduled_at_implies_in_served_at ⇒ //.
+ rewrite another_hep_job_diff_task // same_task_sym.
exact: (diff_task tsk).
Qed.
End FromDifferentTask.
task_interference arr_seq sched j t.
Proof.
apply/andP; split.
- rewrite /non_self task_served_eq_task_scheduled ⇒ //.
apply: job_of_other_task_scheduled ⇒ //.
by move: H_j_tsk ⇒ /eqP → .
- apply/orP; right.
apply/hasP; ∃ j'.
+ by apply scheduled_at_implies_in_served_at ⇒ //.
+ rewrite another_hep_job_diff_task // same_task_sym.
exact: (diff_task tsk).
Qed.
End FromDifferentTask.
Variable j' : Job.
Hypothesis H_j'_sched : scheduled_at sched j' t.
Hypothesis H_j'_lp : ~~ hep_job j' j.
Lemma sched_alp_implies_interference_ahep_false :
¬ another_hep_job_interference arr_seq sched j t.
Proof.
move/hasP ⇒ [jlp /[!mem_filter]/andP[+ IN] AHEP].
move/service_at_implies_scheduled_at ⇒ RSERV.
have EQ := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j'_sched RSERV; subst j'.
by move: (H_j'_lp) AHEP ⇒ LP /andP [HEP A]; rewrite HEP in LP.
Qed.
End LowerPriority.
End ScheduledJob.
Hypothesis H_j'_sched : scheduled_at sched j' t.
Hypothesis H_j'_lp : ~~ hep_job j' j.
Lemma sched_alp_implies_interference_ahep_false :
¬ another_hep_job_interference arr_seq sched j t.
Proof.
move/hasP ⇒ [jlp /[!mem_filter]/andP[+ IN] AHEP].
move/service_at_implies_scheduled_at ⇒ RSERV.
have EQ := ideal_proc_model_is_a_uniprocessor_model _ _ _ _ H_j'_sched RSERV; subst j'.
by move: (H_j'_lp) AHEP ⇒ LP /andP [HEP A]; rewrite HEP in LP.
Qed.
End LowerPriority.
End ScheduledJob.
We prove that we can split cumulative interference into two
parts: (1) cumulative priority inversion and (2) cumulative
interference from jobs with higher or equal priority.
Lemma cumulative_interference_split :
∀ j t1 t2,
cumulative_interference j t1 t2
= cumulative_priority_inversion arr_seq sched j t1 t2
+ cumulative_another_hep_job_interference arr_seq sched j t1 t2.
Proof.
rewrite /cumulative_interference /cumul_cond_interference /cond_interference /interference.
move⇒ j t1 t2; rewrite -big_split //=.
apply/eqP; rewrite eqn_leq; apply/andP; split; rewrite leq_sum//.
{ move ⇒ t _; unfold ideal_jlfp_interference.
by destruct (priority_inversion _ _ _ _), (another_hep_job_interference arr_seq sched j t).
}
{ move ⇒ t _; rewrite /ideal_jlfp_interference.
destruct (ideal_proc_model_sched_case_analysis sched t) as [IDLE | [s SCHED]].
- have → : priority_inversion arr_seq sched j t = false.
exact/negbTE/idle_implies_no_priority_inversion.
by rewrite add0n.
- destruct (hep_job s j) eqn:PRIO.
+ by erewrite sched_hep_implies_no_priority_inversion ⇒ //; rewrite add0n.
+ erewrite !sched_lp_implies_priority_inversion ⇒ //; last by rewrite PRIO.
rewrite orTb.
destruct (another_hep_job_interference arr_seq sched j t) eqn:IAHEP; [exfalso | by []].
eapply sched_alp_implies_interference_ahep_false ⇒ //.
by rewrite PRIO.
}
Qed.
∀ j t1 t2,
cumulative_interference j t1 t2
= cumulative_priority_inversion arr_seq sched j t1 t2
+ cumulative_another_hep_job_interference arr_seq sched j t1 t2.
Proof.
rewrite /cumulative_interference /cumul_cond_interference /cond_interference /interference.
move⇒ j t1 t2; rewrite -big_split //=.
apply/eqP; rewrite eqn_leq; apply/andP; split; rewrite leq_sum//.
{ move ⇒ t _; unfold ideal_jlfp_interference.
by destruct (priority_inversion _ _ _ _), (another_hep_job_interference arr_seq sched j t).
}
{ move ⇒ t _; rewrite /ideal_jlfp_interference.
destruct (ideal_proc_model_sched_case_analysis sched t) as [IDLE | [s SCHED]].
- have → : priority_inversion arr_seq sched j t = false.
exact/negbTE/idle_implies_no_priority_inversion.
by rewrite add0n.
- destruct (hep_job s j) eqn:PRIO.
+ by erewrite sched_hep_implies_no_priority_inversion ⇒ //; rewrite add0n.
+ erewrite !sched_lp_implies_priority_inversion ⇒ //; last by rewrite PRIO.
rewrite orTb.
destruct (another_hep_job_interference arr_seq sched j t) eqn:IAHEP; [exfalso | by []].
eapply sched_alp_implies_interference_ahep_false ⇒ //.
by rewrite PRIO.
}
Qed.
Similarly, we prove that we can split cumulative interfering
workload into two parts: (1) cumulative priority inversion and
(2) cumulative interfering workload from jobs with higher or
equal priority.
Lemma cumulative_interfering_workload_split :
∀ j t1 t2,
cumulative_interfering_workload j t1 t2 =
cumulative_priority_inversion arr_seq sched j t1 t2
+ cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2.
Proof.
rewrite /cumulative_interfering_workload
/cumulative_priority_inversion
/cumulative_other_hep_jobs_interfering_workload.
by move ⇒ j t1 t2; rewrite -big_split //=.
Qed.
∀ j t1 t2,
cumulative_interfering_workload j t1 t2 =
cumulative_priority_inversion arr_seq sched j t1 t2
+ cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2.
Proof.
rewrite /cumulative_interfering_workload
/cumulative_priority_inversion
/cumulative_other_hep_jobs_interfering_workload.
by move ⇒ j t1 t2; rewrite -big_split //=.
Qed.
Before we prove a lemma about the task's interference split,
we show that any job j of task tsk experiences either
priority inversion or task interference if two properties are
satisfied: (1) task tsk is not scheduled at a time instant
t and (2) there is a job jo that experiences interference
at a time t.
Remark priority_inversion_xor_atask_hep_job_interference :
∀ j t,
job_of_task tsk j →
¬ task_scheduled_at arr_seq sched tsk t →
∀ jo,
interference jo t →
(~~ priority_inversion arr_seq sched j t && another_task_hep_job_interference arr_seq sched j t)
|| (priority_inversion arr_seq sched j t && ~~ another_task_hep_job_interference arr_seq sched j t).
Proof.
move ⇒ j t TSK TNSCHED jo INT.
have [PI | PI] /= := boolP (priority_inversion arr_seq sched j t).
- apply/hasP ⇒ -[jhp /[!mem_filter]/andP[RS__jhp IN__jhp]].
move⇒ /andP[ATHEP__hp BB].
have [jlp SCHED__jlp LEP__jlp] :
exists2 j', scheduled_at sched j' t & ~~ hep_job j' j.
exact/uni_priority_inversion_P.
enough (EQ: jlp = jhp); first by (subst; rewrite ATHEP__hp in LEP__jlp).
apply: ideal_proc_model_is_a_uniprocessor_model ⇒ //; move: RS__jhp.
by rewrite /receives_service_at service_at_is_scheduled_at lt0b.
- destruct another_task_hep_job_interference eqn:IATHEP ⇒ //; exfalso.
have L1: interference jo t → ∃ jt, scheduled_at sched jt t.
{ clear PI IATHEP; move ⇒ /orP[].
{ exact: priority_inversion_scheduled_at. }
{ move⇒ /hasP[jt /[!mem_filter]/andP[RSERV _] _].
by ∃ jt; move: RSERV; rewrite /receives_service_at service_at_is_scheduled_at lt0b. }
}
apply L1 in INT; destruct INT as [jt SCHED]; clear L1.
apply: (negP PI).
move: IATHEP ⇒ /eqP; rewrite eqbF_neg.
rewrite /another_task_hep_job_interference.
rewrite has_filter -filter_predI -has_filter ⇒ /hasPn OH.
specialize (OH jt); feed OH.
{ apply arrived_between_implies_in_arrivals ⇒ //.
by apply/andP; split; last (rewrite ltnS //; apply H_jobs_must_arrive_to_execute).
}
rewrite //= negb_and negb_and in OH.
move: OH ⇒ /orP [/orP [OH11 | OH22] | OH2].
+ exact/uni_priority_inversion_P.
+ exfalso; apply: TNSCHED.
apply: (job_of_task_scheduled _ _ _ _ _ _ _ jt) ⇒ //.
move: OH22; rewrite negbK ⇒ /eqP EQ.
by rewrite /job_of_task EQ.
+ exfalso; move: OH2 ⇒ /negP OH2; apply: OH2.
by rewrite /receives_service_at service_at_is_scheduled_at lt0b.
Qed.
∀ j t,
job_of_task tsk j →
¬ task_scheduled_at arr_seq sched tsk t →
∀ jo,
interference jo t →
(~~ priority_inversion arr_seq sched j t && another_task_hep_job_interference arr_seq sched j t)
|| (priority_inversion arr_seq sched j t && ~~ another_task_hep_job_interference arr_seq sched j t).
Proof.
move ⇒ j t TSK TNSCHED jo INT.
have [PI | PI] /= := boolP (priority_inversion arr_seq sched j t).
- apply/hasP ⇒ -[jhp /[!mem_filter]/andP[RS__jhp IN__jhp]].
move⇒ /andP[ATHEP__hp BB].
have [jlp SCHED__jlp LEP__jlp] :
exists2 j', scheduled_at sched j' t & ~~ hep_job j' j.
exact/uni_priority_inversion_P.
enough (EQ: jlp = jhp); first by (subst; rewrite ATHEP__hp in LEP__jlp).
apply: ideal_proc_model_is_a_uniprocessor_model ⇒ //; move: RS__jhp.
by rewrite /receives_service_at service_at_is_scheduled_at lt0b.
- destruct another_task_hep_job_interference eqn:IATHEP ⇒ //; exfalso.
have L1: interference jo t → ∃ jt, scheduled_at sched jt t.
{ clear PI IATHEP; move ⇒ /orP[].
{ exact: priority_inversion_scheduled_at. }
{ move⇒ /hasP[jt /[!mem_filter]/andP[RSERV _] _].
by ∃ jt; move: RSERV; rewrite /receives_service_at service_at_is_scheduled_at lt0b. }
}
apply L1 in INT; destruct INT as [jt SCHED]; clear L1.
apply: (negP PI).
move: IATHEP ⇒ /eqP; rewrite eqbF_neg.
rewrite /another_task_hep_job_interference.
rewrite has_filter -filter_predI -has_filter ⇒ /hasPn OH.
specialize (OH jt); feed OH.
{ apply arrived_between_implies_in_arrivals ⇒ //.
by apply/andP; split; last (rewrite ltnS //; apply H_jobs_must_arrive_to_execute).
}
rewrite //= negb_and negb_and in OH.
move: OH ⇒ /orP [/orP [OH11 | OH22] | OH2].
+ exact/uni_priority_inversion_P.
+ exfalso; apply: TNSCHED.
apply: (job_of_task_scheduled _ _ _ _ _ _ _ jt) ⇒ //.
move: OH22; rewrite negbK ⇒ /eqP EQ.
by rewrite /job_of_task EQ.
+ exfalso; move: OH2 ⇒ /negP OH2; apply: OH2.
by rewrite /receives_service_at service_at_is_scheduled_at lt0b.
Qed.
Let j be any job of task tsk. Then the cumulative task
interference received by job j is bounded to the sum of the
cumulative priority inversion of job j and the cumulative
interference incurred by job j due to higher-or-equal
priority jobs from other tasks.
Lemma cumulative_task_interference_split :
∀ j t1 t2,
arrives_in arr_seq j →
job_of_task tsk j →
~~ completed_by sched j t2 →
cumul_task_interference arr_seq sched j t1 t2
≤ cumulative_priority_inversion arr_seq sched j t1 t2
+ cumulative_another_task_hep_job_interference arr_seq sched j t1 t2.
Proof.
move⇒ j t1 R ARR TSK NCOMPL.
rewrite /cumul_task_interference /cumul_cond_interference.
rewrite -big_split //= big_seq_cond [leqRHS]big_seq_cond.
apply leq_sum; move ⇒ t /andP [IN _].
rewrite /cond_interference /non_self /interference /ideal_jlfp_interference.
have [IDLE|[s SCHEDs]] := ideal_proc_model_sched_case_analysis sched t.
{ move: (IDLE) ⇒ IIDLE; erewrite <-is_idle_def in IDLE ⇒ //.
have ->: priority_inversion arr_seq sched j t = false
by apply/eqP; rewrite eqbF_neg; apply: no_priority_inversion_when_idle ⇒ //.
have ->: another_hep_job_interference arr_seq sched j t = false
by apply/eqP; rewrite eqbF_neg; apply/negP; apply idle_implies_no_hep_job_interference.
have ->: another_task_hep_job_interference arr_seq sched j t = false
by apply/eqP; rewrite eqbF_neg; apply/negP; apply idle_implies_no_hep_task_interference.
by rewrite andbF.
}
{ rewrite (interference_ahep_equiv_ahep _ _ _ SCHEDs).
rewrite (interference_athep_equiv_athep _ _ _ SCHEDs) ⇒ //.
rewrite (priority_inversion_equiv_sched_lower_priority _ _ _ _ _ _ _ _ _ SCHEDs) ⇒ //.
rewrite task_served_eq_task_scheduled // (job_of_scheduled_task _ _ _ _ _ _ _ _ _ SCHEDs) ⇒ //.
rewrite /another_hep_job /another_task_hep_job.
have [EQj|NEQj] := eqVneq s j.
{ by subst; rewrite /job_of_task eq_refl H_priority_is_reflexive. }
have [/eqP EQt|NEQt] := eqVneq (job_task s) (job_task j).
{ apply/eqP; move: (EQt) ⇒ /eqP <-.
by rewrite /job_of_task eq_refl //= andbF addn0 eq_sym eqb0; apply/negP ⇒ LPs. }
{ by rewrite /job_of_task NEQt //= andbT; case: hep_job. }
}
Qed.
∀ j t1 t2,
arrives_in arr_seq j →
job_of_task tsk j →
~~ completed_by sched j t2 →
cumul_task_interference arr_seq sched j t1 t2
≤ cumulative_priority_inversion arr_seq sched j t1 t2
+ cumulative_another_task_hep_job_interference arr_seq sched j t1 t2.
Proof.
move⇒ j t1 R ARR TSK NCOMPL.
rewrite /cumul_task_interference /cumul_cond_interference.
rewrite -big_split //= big_seq_cond [leqRHS]big_seq_cond.
apply leq_sum; move ⇒ t /andP [IN _].
rewrite /cond_interference /non_self /interference /ideal_jlfp_interference.
have [IDLE|[s SCHEDs]] := ideal_proc_model_sched_case_analysis sched t.
{ move: (IDLE) ⇒ IIDLE; erewrite <-is_idle_def in IDLE ⇒ //.
have ->: priority_inversion arr_seq sched j t = false
by apply/eqP; rewrite eqbF_neg; apply: no_priority_inversion_when_idle ⇒ //.
have ->: another_hep_job_interference arr_seq sched j t = false
by apply/eqP; rewrite eqbF_neg; apply/negP; apply idle_implies_no_hep_job_interference.
have ->: another_task_hep_job_interference arr_seq sched j t = false
by apply/eqP; rewrite eqbF_neg; apply/negP; apply idle_implies_no_hep_task_interference.
by rewrite andbF.
}
{ rewrite (interference_ahep_equiv_ahep _ _ _ SCHEDs).
rewrite (interference_athep_equiv_athep _ _ _ SCHEDs) ⇒ //.
rewrite (priority_inversion_equiv_sched_lower_priority _ _ _ _ _ _ _ _ _ SCHEDs) ⇒ //.
rewrite task_served_eq_task_scheduled // (job_of_scheduled_task _ _ _ _ _ _ _ _ _ SCHEDs) ⇒ //.
rewrite /another_hep_job /another_task_hep_job.
have [EQj|NEQj] := eqVneq s j.
{ by subst; rewrite /job_of_task eq_refl H_priority_is_reflexive. }
have [/eqP EQt|NEQt] := eqVneq (job_task s) (job_task j).
{ apply/eqP; move: (EQt) ⇒ /eqP <-.
by rewrite /job_of_task eq_refl //= andbF addn0 eq_sym eqb0; apply/negP ⇒ LPs. }
{ by rewrite /job_of_task NEQt //= andbT; case: hep_job. }
}
Qed.
In this section, we prove that the (abstract) cumulative
interfering workload is equivalent to the conventional workload,
i.e., the one defined with concrete schedule parameters.
Let
[t1,t2)
be any time interval.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Then for any job j, the cumulative interfering workload is
equal to the conventional workload.
Lemma cumulative_iw_hep_eq_workload_of_ohep :
cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2
= workload_of_another_hep_jobs j t1 t2.
Proof.
rewrite /cumulative_other_hep_jobs_interfering_workload /workload_of_another_hep_jobs.
case NEQ: (t1 < t2); last first.
{ move: NEQ ⇒ /negP /negP; rewrite -leqNgt; move ⇒ NEQ.
rewrite big_geq//.
rewrite /arrivals_between /arrival_sequence.arrivals_between big_geq//.
by rewrite /workload_of_jobs big_nil.
}
{ unfold other_hep_jobs_interfering_workload, workload_of_jobs.
interval_to_duration t1 t2 k.
elim: k ⇒ [|k IHk].
- rewrite !addn0.
rewrite big_geq//.
rewrite /arrivals_between /arrival_sequence.arrivals_between big_geq//.
by rewrite /workload_of_jobs big_nil.
- rewrite addnS big_nat_recr //=; last by rewrite leq_addr.
rewrite IHk.
rewrite /arrivals_between /arrival_sequence.arrivals_between big_nat_recr //=.
+ by rewrite big_cat.
+ by rewrite leq_addr.
}
Qed.
End InstantiatedWorkloadEquivalence.
cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2
= workload_of_another_hep_jobs j t1 t2.
Proof.
rewrite /cumulative_other_hep_jobs_interfering_workload /workload_of_another_hep_jobs.
case NEQ: (t1 < t2); last first.
{ move: NEQ ⇒ /negP /negP; rewrite -leqNgt; move ⇒ NEQ.
rewrite big_geq//.
rewrite /arrivals_between /arrival_sequence.arrivals_between big_geq//.
by rewrite /workload_of_jobs big_nil.
}
{ unfold other_hep_jobs_interfering_workload, workload_of_jobs.
interval_to_duration t1 t2 k.
elim: k ⇒ [|k IHk].
- rewrite !addn0.
rewrite big_geq//.
rewrite /arrivals_between /arrival_sequence.arrivals_between big_geq//.
by rewrite /workload_of_jobs big_nil.
- rewrite addnS big_nat_recr //=; last by rewrite leq_addr.
rewrite IHk.
rewrite /arrivals_between /arrival_sequence.arrivals_between big_nat_recr //=.
+ by rewrite big_cat.
+ by rewrite leq_addr.
}
Qed.
End InstantiatedWorkloadEquivalence.
In this section, we prove that the (abstract) cumulative
interference of jobs with higher or equal priority is equal to
total service of jobs with higher or equal priority.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
We consider an arbitrary time interval
[t1, t)
that
starts with a quiet time.
Then for job j, the (abstract) instantiated function of
interference is equal to the total service of any subset of jobs
with higher or equal priority.
Lemma cumulative_pred_eq_service (P : pred Job) :
(∀ j', P j' → hep_job j' j) →
\sum_(t1 ≤ t' < t) has P (served_jobs_at arr_seq sched t')
= service_of_jobs sched P (arrivals_between arr_seq t1 t) t1 t.
Proof.
move⇒ Phep; clear H_job_of_tsk.
rewrite [RHS]exchange_big /=; apply: eq_big_nat ⇒ x /andP[t1lex xltt].
have [Idle|[jo Sched_jo]] := ideal_proc_model_sched_case_analysis sched x.
{ rewrite has_filter -filter_predI -has_filter.
rewrite (@eq_has _ _ pred0) ?has_pred0 ?big1 // ⇒ [j' _ | j'].
+ exact: ideal_not_idle_implies_sched.
+ by rewrite /= /receives_service_at ideal_not_idle_implies_sched // andbF. }
have arr_jo : arrives_in arr_seq jo.
{ exact: H_jobs_come_from_arrival_sequence Sched_jo. }
have arr_jo_x : job_arrival jo ≤ x.
{ exact: H_jobs_must_arrive_to_execute Sched_jo. }
have [jo_hep_j|PRIO] := boolP (P jo).
- transitivity true.
{ rewrite has_filter -filter_predI -has_filter.
congr nat_of_bool; apply/hasP; ∃ jo.
- exact: arrived_between_implies_in_arrivals.
- rewrite /= jo_hep_j /receives_service_at service_at_is_scheduled_at.
by rewrite Sched_jo. }
apply/esym/eqP; rewrite eqn_leq; apply/andP; split.
+ exact: service_of_jobs_le_1.
+ rewrite sum_nat_gt0; apply/hasP; ∃ jo; last first.
{ by rewrite service_at_is_scheduled_at Sched_jo. }
rewrite mem_filter jo_hep_j /=.
apply: arrived_between_implies_in_arrivals ⇒ //.
apply/negPn; rewrite negb_and -ltnNge -leqNgt.
apply/negP ⇒ /orP[arr_jo_t1|]; last by lia.
move: Sched_jo; apply/negP.
apply: completed_implies_not_scheduled ⇒ //.
apply: completion_monotonic t1lex _.
by apply: H_quiet_time ⇒ //; apply: Phep jo_hep_j.
- rewrite has_filter -filter_predI -has_filter.
rewrite (@eq_in_has _ _ pred0) ?has_pred0 ?big1// ⇒ [j' AHEP|j' IN].
+ apply/eqP; rewrite service_at_is_scheduled_at eqb0; apply/negP.
move⇒ SCHED; move: AHEP PRIO; suff → : jo = j' by move⇒ →.
exact: ideal_proc_model_is_a_uniprocessor_model.
+ apply/negbTE; rewrite negb_and orbC -implyNb negbK; apply/implyP.
rewrite /receives_service_at service_at_is_scheduled_at lt0b.
move⇒ SCHED; suff <- : jo = j' by [].
exact: ideal_proc_model_is_a_uniprocessor_model.
Qed.
(∀ j', P j' → hep_job j' j) →
\sum_(t1 ≤ t' < t) has P (served_jobs_at arr_seq sched t')
= service_of_jobs sched P (arrivals_between arr_seq t1 t) t1 t.
Proof.
move⇒ Phep; clear H_job_of_tsk.
rewrite [RHS]exchange_big /=; apply: eq_big_nat ⇒ x /andP[t1lex xltt].
have [Idle|[jo Sched_jo]] := ideal_proc_model_sched_case_analysis sched x.
{ rewrite has_filter -filter_predI -has_filter.
rewrite (@eq_has _ _ pred0) ?has_pred0 ?big1 // ⇒ [j' _ | j'].
+ exact: ideal_not_idle_implies_sched.
+ by rewrite /= /receives_service_at ideal_not_idle_implies_sched // andbF. }
have arr_jo : arrives_in arr_seq jo.
{ exact: H_jobs_come_from_arrival_sequence Sched_jo. }
have arr_jo_x : job_arrival jo ≤ x.
{ exact: H_jobs_must_arrive_to_execute Sched_jo. }
have [jo_hep_j|PRIO] := boolP (P jo).
- transitivity true.
{ rewrite has_filter -filter_predI -has_filter.
congr nat_of_bool; apply/hasP; ∃ jo.
- exact: arrived_between_implies_in_arrivals.
- rewrite /= jo_hep_j /receives_service_at service_at_is_scheduled_at.
by rewrite Sched_jo. }
apply/esym/eqP; rewrite eqn_leq; apply/andP; split.
+ exact: service_of_jobs_le_1.
+ rewrite sum_nat_gt0; apply/hasP; ∃ jo; last first.
{ by rewrite service_at_is_scheduled_at Sched_jo. }
rewrite mem_filter jo_hep_j /=.
apply: arrived_between_implies_in_arrivals ⇒ //.
apply/negPn; rewrite negb_and -ltnNge -leqNgt.
apply/negP ⇒ /orP[arr_jo_t1|]; last by lia.
move: Sched_jo; apply/negP.
apply: completed_implies_not_scheduled ⇒ //.
apply: completion_monotonic t1lex _.
by apply: H_quiet_time ⇒ //; apply: Phep jo_hep_j.
- rewrite has_filter -filter_predI -has_filter.
rewrite (@eq_in_has _ _ pred0) ?has_pred0 ?big1// ⇒ [j' AHEP|j' IN].
+ apply/eqP; rewrite service_at_is_scheduled_at eqb0; apply/negP.
move⇒ SCHED; move: AHEP PRIO; suff → : jo = j' by move⇒ →.
exact: ideal_proc_model_is_a_uniprocessor_model.
+ apply/negbTE; rewrite negb_and orbC -implyNb negbK; apply/implyP.
rewrite /receives_service_at service_at_is_scheduled_at lt0b.
move⇒ SCHED; suff <- : jo = j' by [].
exact: ideal_proc_model_is_a_uniprocessor_model.
Qed.
The above is in particular true for the jobs other
than j with higher or equal priority...
Lemma cumulative_i_ohep_eq_service_of_ohep :
cumulative_another_hep_job_interference arr_seq sched j t1 t
= service_of_another_hep_jobs j t1 t.
Proof. by apply: cumulative_pred_eq_service ⇒ ? /andP[]. Qed.
cumulative_another_hep_job_interference arr_seq sched j t1 t
= service_of_another_hep_jobs j t1 t.
Proof. by apply: cumulative_pred_eq_service ⇒ ? /andP[]. Qed.
...and for jobs from other tasks than j with higher
or equal priority.
Lemma cumulative_i_thep_eq_service_of_othep :
cumulative_another_task_hep_job_interference arr_seq sched j t1 t
= service_of_another_task_hep_job j t1 t.
Proof. by apply: cumulative_pred_eq_service ⇒ ? /andP[]. Qed.
End InstantiatedServiceEquivalences.
cumulative_another_task_hep_job_interference arr_seq sched j t1 t
= service_of_another_task_hep_job j t1 t.
Proof. by apply: cumulative_pred_eq_service ⇒ ? /andP[]. Qed.
End InstantiatedServiceEquivalences.
In this section we prove that the abstract definition of busy
interval is equivalent to the conventional, concrete
definition of busy interval for JLFP scheduling.
In order to avoid confusion, we denote the notion of a quiet
time in the classical sense as quiet_time_cl, and the
notion of quiet time in the abstract sense as
quiet_time_ab.
Let quiet_time_cl := busy_interval.quiet_time arr_seq sched.
Let quiet_time_ab := definitions.quiet_time sched.
Let quiet_time_ab := definitions.quiet_time sched.
Same for the two notions of a busy interval prefix ...
Let busy_interval_prefix_cl := busy_interval.busy_interval_prefix arr_seq sched.
Let busy_interval_prefix_ab := definitions.busy_interval_prefix sched.
Let busy_interval_prefix_ab := definitions.busy_interval_prefix sched.
... and the two notions of a busy interval.
Let busy_interval_cl := busy_interval.busy_interval arr_seq sched.
Let busy_interval_ab := definitions.busy_interval sched.
Let busy_interval_ab := definitions.busy_interval sched.
Consider any job j of tsk.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_cost_positive : job_cost_positive j.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_cost_positive : job_cost_positive j.
To show the equivalence of the notions of busy intervals, we
first show that the notions of quiet time are also
equivalent.
First, we show that the classical notion of quiet time
implies the abstract notion of quiet time.
Lemma quiet_time_cl_implies_quiet_time_ab :
∀ t, quiet_time_cl j t → quiet_time_ab j t.
Proof.
clear H_JLFP_respects_sequential_tasks.
have zero_is_quiet_time: ∀ j, quiet_time_cl j 0.
{ by move ⇒ jhp ARR HP AB; move: AB; rewrite /arrived_before ltn0. }
move⇒ t QT; apply/andP; split; last first.
{ rewrite negb_and negbK; apply/orP.
by case ARR: (arrived_before j t); [right | left]; [apply QT | ]. }
{ erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l.
rewrite cumulative_i_ohep_eq_service_of_ohep//.
rewrite //= cumulative_iw_hep_eq_workload_of_ohep eq_sym; apply/eqP.
apply all_jobs_have_completed_equiv_workload_eq_service ⇒ //.
move ⇒ j0 IN HEP; apply QT.
- by apply in_arrivals_implies_arrived in IN.
- by move: HEP ⇒ /andP [H6 H7].
- by apply in_arrivals_implies_arrived_between in IN.
}
Qed.
∀ t, quiet_time_cl j t → quiet_time_ab j t.
Proof.
clear H_JLFP_respects_sequential_tasks.
have zero_is_quiet_time: ∀ j, quiet_time_cl j 0.
{ by move ⇒ jhp ARR HP AB; move: AB; rewrite /arrived_before ltn0. }
move⇒ t QT; apply/andP; split; last first.
{ rewrite negb_and negbK; apply/orP.
by case ARR: (arrived_before j t); [right | left]; [apply QT | ]. }
{ erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l.
rewrite cumulative_i_ohep_eq_service_of_ohep//.
rewrite //= cumulative_iw_hep_eq_workload_of_ohep eq_sym; apply/eqP.
apply all_jobs_have_completed_equiv_workload_eq_service ⇒ //.
move ⇒ j0 IN HEP; apply QT.
- by apply in_arrivals_implies_arrived in IN.
- by move: HEP ⇒ /andP [H6 H7].
- by apply in_arrivals_implies_arrived_between in IN.
}
Qed.
And vice versa, the abstract notion of quiet time implies
the classical notion of quiet time.
Lemma quiet_time_ab_implies_quiet_time_cl :
∀ t, quiet_time_ab j t → quiet_time_cl j t.
Proof.
clear H_JLFP_respects_sequential_tasks.
have zero_is_quiet_time: ∀ j, quiet_time_cl j 0.
{ by move ⇒ jhp ARR HP AB; move: AB; rewrite /arrived_before ltn0. }
move ⇒ t /andP [T0 T1] jhp ARR HP ARB.
eapply all_jobs_have_completed_equiv_workload_eq_service with
(P := fun jhp ⇒ hep_job jhp j) (t1 := 0) (t2 := t) ⇒ //.
erewrite service_of_jobs_case_on_pred with (P2 := fun j' ⇒ j' != j).
erewrite workload_of_jobs_case_on_pred with (P' := fun j' ⇒ j' != j) ⇒ //.
replace ((fun j0 : Job ⇒ hep_job j0 j && (j0 != j))) with (another_hep_job^~j); last by rewrite /another_hep_job.
rewrite -/(service_of_another_hep_jobs j 0 t) -cumulative_i_ohep_eq_service_of_ohep; eauto.
rewrite -/(workload_of_another_hep_jobs j 0 t) -cumulative_iw_hep_eq_workload_of_ohep; eauto.
move: T1; rewrite negb_and ⇒ /orP [NA | /negPn COMP].
{ have PRED__eq: {in arrivals_between arr_seq 0 t, (fun j__copy : Job ⇒ hep_job j__copy j && ~~ (j__copy != j)) =1 pred0}.
{ move ⇒ j__copy IN; apply negbTE.
rewrite negb_and; apply/orP; right; apply/negPn/eqP ⇒ EQ; subst j__copy.
move: NA ⇒ /negP CONTR; apply: CONTR.
by apply in_arrivals_implies_arrived_between in IN. }
erewrite service_of_jobs_equiv_pred with (P2 := pred0) ⇒ [|//].
erewrite workload_of_jobs_equiv_pred with (P' := pred0) ⇒ [|//].
move: T0; erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l ⇒ /eqP EQ.
rewrite EQ; clear EQ; apply/eqP; rewrite eqn_add2l.
by erewrite workload_of_jobs_pred0, service_of_jobs_pred0.
}
{ have PRED__eq: {in arrivals_between arr_seq 0 t, (fun j0 : Job ⇒ hep_job j0 j && ~~ (j0 != j)) =1 eq_op j}.
{ move ⇒ j__copy IN.
replace (~~ (j__copy != j)) with (j__copy == j); last by case: (j__copy == j).
rewrite eq_sym; destruct (j == j__copy) eqn:EQ; last by rewrite Bool.andb_false_r.
by move: EQ ⇒ /eqP EQ; rewrite Bool.andb_true_r; apply/eqP; rewrite eqb_id; subst. }
erewrite service_of_jobs_equiv_pred with (P2 := eq_op j) ⇒ [|//].
erewrite workload_of_jobs_equiv_pred with (P' := eq_op j) ⇒ [|//].
move: T0; erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l ⇒ /eqP EQ.
rewrite EQ; clear EQ; apply/eqP; rewrite eqn_add2l.
apply/eqP; eapply all_jobs_have_completed_equiv_workload_eq_service with
(P := eq_op j) (t1 := 0) (t2 := t) ⇒ //.
by move ⇒ j__copy _ /eqP EQ; subst j__copy.
}
move ⇒ ? _ _ ; unfold arrived_before; lia.
Qed.
∀ t, quiet_time_ab j t → quiet_time_cl j t.
Proof.
clear H_JLFP_respects_sequential_tasks.
have zero_is_quiet_time: ∀ j, quiet_time_cl j 0.
{ by move ⇒ jhp ARR HP AB; move: AB; rewrite /arrived_before ltn0. }
move ⇒ t /andP [T0 T1] jhp ARR HP ARB.
eapply all_jobs_have_completed_equiv_workload_eq_service with
(P := fun jhp ⇒ hep_job jhp j) (t1 := 0) (t2 := t) ⇒ //.
erewrite service_of_jobs_case_on_pred with (P2 := fun j' ⇒ j' != j).
erewrite workload_of_jobs_case_on_pred with (P' := fun j' ⇒ j' != j) ⇒ //.
replace ((fun j0 : Job ⇒ hep_job j0 j && (j0 != j))) with (another_hep_job^~j); last by rewrite /another_hep_job.
rewrite -/(service_of_another_hep_jobs j 0 t) -cumulative_i_ohep_eq_service_of_ohep; eauto.
rewrite -/(workload_of_another_hep_jobs j 0 t) -cumulative_iw_hep_eq_workload_of_ohep; eauto.
move: T1; rewrite negb_and ⇒ /orP [NA | /negPn COMP].
{ have PRED__eq: {in arrivals_between arr_seq 0 t, (fun j__copy : Job ⇒ hep_job j__copy j && ~~ (j__copy != j)) =1 pred0}.
{ move ⇒ j__copy IN; apply negbTE.
rewrite negb_and; apply/orP; right; apply/negPn/eqP ⇒ EQ; subst j__copy.
move: NA ⇒ /negP CONTR; apply: CONTR.
by apply in_arrivals_implies_arrived_between in IN. }
erewrite service_of_jobs_equiv_pred with (P2 := pred0) ⇒ [|//].
erewrite workload_of_jobs_equiv_pred with (P' := pred0) ⇒ [|//].
move: T0; erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l ⇒ /eqP EQ.
rewrite EQ; clear EQ; apply/eqP; rewrite eqn_add2l.
by erewrite workload_of_jobs_pred0, service_of_jobs_pred0.
}
{ have PRED__eq: {in arrivals_between arr_seq 0 t, (fun j0 : Job ⇒ hep_job j0 j && ~~ (j0 != j)) =1 eq_op j}.
{ move ⇒ j__copy IN.
replace (~~ (j__copy != j)) with (j__copy == j); last by case: (j__copy == j).
rewrite eq_sym; destruct (j == j__copy) eqn:EQ; last by rewrite Bool.andb_false_r.
by move: EQ ⇒ /eqP EQ; rewrite Bool.andb_true_r; apply/eqP; rewrite eqb_id; subst. }
erewrite service_of_jobs_equiv_pred with (P2 := eq_op j) ⇒ [|//].
erewrite workload_of_jobs_equiv_pred with (P' := eq_op j) ⇒ [|//].
move: T0; erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l ⇒ /eqP EQ.
rewrite EQ; clear EQ; apply/eqP; rewrite eqn_add2l.
apply/eqP; eapply all_jobs_have_completed_equiv_workload_eq_service with
(P := eq_op j) (t1 := 0) (t2 := t) ⇒ //.
by move ⇒ j__copy _ /eqP EQ; subst j__copy.
}
move ⇒ ? _ _ ; unfold arrived_before; lia.
Qed.
The equivalence trivially follows from the lemmas above.
Corollary instantiated_quiet_time_equivalent_quiet_time :
∀ t,
quiet_time_cl j t ↔ quiet_time_ab j t.
Proof.
clear H_JLFP_respects_sequential_tasks.
move ⇒ ?; split.
- by apply quiet_time_cl_implies_quiet_time_ab.
- by apply quiet_time_ab_implies_quiet_time_cl.
Qed.
∀ t,
quiet_time_cl j t ↔ quiet_time_ab j t.
Proof.
clear H_JLFP_respects_sequential_tasks.
move ⇒ ?; split.
- by apply quiet_time_cl_implies_quiet_time_ab.
- by apply quiet_time_ab_implies_quiet_time_cl.
Qed.
Based on that, we prove that the concept of busy interval
prefix obtained by instantiating the abstract definition of
busy interval prefix coincides with the conventional
definition of busy interval prefix.
Lemma instantiated_busy_interval_prefix_equivalent_busy_interval_prefix :
∀ t1 t2, busy_interval_prefix_cl j t1 t2 ↔ busy_interval_prefix_ab j t1 t2.
Proof.
move ⇒ t1 t2; split.
{ move ⇒ [NEQ [QTt1 [NQT REL]]].
split⇒ [//|]; split.
- by apply instantiated_quiet_time_equivalent_quiet_time in QTt1.
- by move ⇒ t NE QT; eapply NQT; eauto 2; apply instantiated_quiet_time_equivalent_quiet_time.
}
{ move ⇒ [/andP [NEQ1 NEQ2] [QTt1 NQT]].
split; [ | split; [ |split] ].
- by apply leq_ltn_trans with (job_arrival j).
- by eapply instantiated_quiet_time_equivalent_quiet_time.
- by move ⇒ t NEQ QT; eapply NQT; eauto 2; eapply instantiated_quiet_time_equivalent_quiet_time in QT.
- by apply/andP.
}
Qed.
∀ t1 t2, busy_interval_prefix_cl j t1 t2 ↔ busy_interval_prefix_ab j t1 t2.
Proof.
move ⇒ t1 t2; split.
{ move ⇒ [NEQ [QTt1 [NQT REL]]].
split⇒ [//|]; split.
- by apply instantiated_quiet_time_equivalent_quiet_time in QTt1.
- by move ⇒ t NE QT; eapply NQT; eauto 2; apply instantiated_quiet_time_equivalent_quiet_time.
}
{ move ⇒ [/andP [NEQ1 NEQ2] [QTt1 NQT]].
split; [ | split; [ |split] ].
- by apply leq_ltn_trans with (job_arrival j).
- by eapply instantiated_quiet_time_equivalent_quiet_time.
- by move ⇒ t NEQ QT; eapply NQT; eauto 2; eapply instantiated_quiet_time_equivalent_quiet_time in QT.
- by apply/andP.
}
Qed.
Similarly, we prove that the concept of busy interval
obtained by instantiating the abstract definition of busy
interval coincides with the conventional definition of busy
interval.
Lemma instantiated_busy_interval_equivalent_busy_interval :
∀ t1 t2, busy_interval_cl j t1 t2 ↔ busy_interval_ab j t1 t2.
Proof.
move ⇒ t1 t2; split.
{ move ⇒ [PREF QTt2]; split.
- by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- by eapply instantiated_quiet_time_equivalent_quiet_time in QTt2.
}
{ move ⇒ [PREF QTt2]; split.
- by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- by eapply instantiated_quiet_time_equivalent_quiet_time.
}
Qed.
∀ t1 t2, busy_interval_cl j t1 t2 ↔ busy_interval_ab j t1 t2.
Proof.
move ⇒ t1 t2; split.
{ move ⇒ [PREF QTt2]; split.
- by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- by eapply instantiated_quiet_time_equivalent_quiet_time in QTt2.
}
{ move ⇒ [PREF QTt2]; split.
- by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- by eapply instantiated_quiet_time_equivalent_quiet_time.
}
Qed.
For the sake of proof automation, we note the frequently needed
special case of an abstract busy window implying the existence of a
classic quiet time.
Fact abstract_busy_interval_classic_quiet_time :
∀ t1 t2,
busy_interval_ab j t1 t2 → quiet_time_cl j t1.
Proof.
by move⇒ ? ? /instantiated_busy_interval_equivalent_busy_interval [[_ [? _]] _].
Qed.
∀ t1 t2,
busy_interval_ab j t1 t2 → quiet_time_cl j t1.
Proof.
by move⇒ ? ? /instantiated_busy_interval_equivalent_busy_interval [[_ [? _]] _].
Qed.
Also for automation, we note a similar fact about classic busy-window prefixes.
Fact abstract_busy_interval_classic_busy_interval_prefix :
∀ t1 t2,
busy_interval_ab j t1 t2 → busy_interval_prefix_cl j t1 t2.
Proof. by move⇒ ? ? /instantiated_busy_interval_equivalent_busy_interval [+ _]. Qed.
End BusyIntervalEquivalence.
End Equivalences.
∀ t1 t2,
busy_interval_ab j t1 t2 → busy_interval_prefix_cl j t1 t2.
Proof. by move⇒ ? ? /instantiated_busy_interval_equivalent_busy_interval [+ _]. Qed.
End BusyIntervalEquivalence.
End Equivalences.
In this section we prove some properties about the interference
and interfering workload as defined in this file.
Consider work-bearing readiness.
Context `{@JobReady Job (ideal.processor_state Job) _ _}.
Hypothesis H_work_bearing_readiness : work_bearing_readiness arr_seq sched.
Hypothesis H_work_bearing_readiness : work_bearing_readiness arr_seq sched.
Assume that the schedule is valid and work-conserving.
Note that we differentiate between abstract and classical
notions of work-conserving schedule.
Let work_conserving_ab := definitions.work_conserving arr_seq sched.
Let work_conserving_cl := work_conserving.work_conserving arr_seq sched.
Let busy_interval_prefix_ab := definitions.busy_interval_prefix sched.
Let work_conserving_cl := work_conserving.work_conserving arr_seq sched.
Let busy_interval_prefix_ab := definitions.busy_interval_prefix sched.
We assume that the schedule is a work-conserving schedule in
the classical sense, and later prove that the hypothesis
about abstract work-conservation also holds.
Assume the scheduling policy under consideration is reflexive.
In this section, we prove the correctness of interference
inside the busy interval, i.e., we prove that if interference
for a job is false then the job is scheduled and vice versa.
This property is referred to as abstract work conservation.
Consider a job j that is in the arrival sequence
and has a positive job cost.
Variable j : Job.
Hypothesis H_arrives : arrives_in arr_seq j.
Hypothesis H_job_cost_positive : 0 < job_cost j.
Hypothesis H_arrives : arrives_in arr_seq j.
Hypothesis H_job_cost_positive : 0 < job_cost j.
Let the busy interval of the job be
[t1,t2)
.
Consider a time t inside the busy interval of the job.
Lemma not_interference_implies_scheduled :
¬ interference j t → receives_service_at sched j t.
Proof.
move ⇒ /negP HYP; move : HYP.
rewrite negb_or /another_hep_job_interference.
move ⇒ /andP [HYP1 HYP2].
have [Idle|[jo Sched_jo]] := ideal_proc_model_sched_case_analysis sched t.
{ exfalso; clear HYP1 HYP2.
eapply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix in H_busy_interval_prefix ⇒ //.
apply: not_quiet_implies_not_idle ⇒ //.
by rewrite is_idle_def.
}
{ have PINV := negP HYP1.
have: ~~ another_hep_job jo j.
{ apply: contraNN HYP2.
by rewrite -(interference_ahep_equiv_ahep _ t). }
rewrite negb_and ⇒ /orP [NHEP | EQ].
- rewrite/receives_service_at; move_neq_up ZS; move: ZS; rewrite leqn0 ⇒ /eqP ZS.
apply no_service_not_scheduled in ZS ⇒ //; apply: PINV.
exact/uni_priority_inversion_P.
- apply negbNE in EQ; move: EQ ⇒ /eqP EQ; subst jo.
by rewrite /receives_service_at service_at_is_scheduled_at Sched_jo.
}
Qed.
¬ interference j t → receives_service_at sched j t.
Proof.
move ⇒ /negP HYP; move : HYP.
rewrite negb_or /another_hep_job_interference.
move ⇒ /andP [HYP1 HYP2].
have [Idle|[jo Sched_jo]] := ideal_proc_model_sched_case_analysis sched t.
{ exfalso; clear HYP1 HYP2.
eapply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix in H_busy_interval_prefix ⇒ //.
apply: not_quiet_implies_not_idle ⇒ //.
by rewrite is_idle_def.
}
{ have PINV := negP HYP1.
have: ~~ another_hep_job jo j.
{ apply: contraNN HYP2.
by rewrite -(interference_ahep_equiv_ahep _ t). }
rewrite negb_and ⇒ /orP [NHEP | EQ].
- rewrite/receives_service_at; move_neq_up ZS; move: ZS; rewrite leqn0 ⇒ /eqP ZS.
apply no_service_not_scheduled in ZS ⇒ //; apply: PINV.
exact/uni_priority_inversion_P.
- apply negbNE in EQ; move: EQ ⇒ /eqP EQ; subst jo.
by rewrite /receives_service_at service_at_is_scheduled_at Sched_jo.
}
Qed.
Lemma scheduled_implies_no_interference :
receives_service_at sched j t → ¬ interference j t.
Proof.
move⇒ RSERV /orP[PINV | INT].
- rewrite (sched_hep_implies_no_priority_inversion _ _ _ _ _ _ _ _ j) in PINV ⇒ //.
by rewrite /receives_service_at service_at_is_scheduled_at lt0b in RSERV.
- rewrite /receives_service_at service_at_is_scheduled_at lt0b in RSERV.
by apply interference_ahep_job_eq_false in RSERV.
Qed.
End Abstract_Work_Conservation.
receives_service_at sched j t → ¬ interference j t.
Proof.
move⇒ RSERV /orP[PINV | INT].
- rewrite (sched_hep_implies_no_priority_inversion _ _ _ _ _ _ _ _ j) in PINV ⇒ //.
by rewrite /receives_service_at service_at_is_scheduled_at lt0b in RSERV.
- rewrite /receives_service_at service_at_is_scheduled_at lt0b in RSERV.
by apply interference_ahep_job_eq_false in RSERV.
Qed.
End Abstract_Work_Conservation.
Using the above two lemmas, we can prove that abstract work
conservation always holds for these instantiations of I and
IW.
Corollary instantiated_i_and_w_are_coherent_with_schedule :
work_conserving_ab.
Proof.
move ⇒ j t1 t2 t ARR POS BUSY NEQ; split.
- exact: (not_interference_implies_scheduled j ARR POS).
- exact: scheduled_implies_no_interference.
Qed.
work_conserving_ab.
Proof.
move ⇒ j t1 t2 t ARR POS BUSY NEQ; split.
- exact: (not_interference_implies_scheduled j ARR POS).
- exact: scheduled_implies_no_interference.
Qed.
Next, in order to prove that these definitions of I and IW
are consistent with sequential tasks, we need to assume that
the policy under consideration respects sequential tasks.
We prove that these definitions of I and IW are consistent
with sequential tasks.
Lemma instantiated_interference_and_workload_consistent_with_sequential_tasks :
interference_and_workload_consistent_with_sequential_tasks arr_seq sched tsk.
Proof.
move ⇒ j t1 t2 ARR /eqP TSK POS BUSY.
eapply instantiated_busy_interval_equivalent_busy_interval in BUSY ⇒ //.
eapply all_jobs_have_completed_equiv_workload_eq_service ⇒ //.
move ⇒ s INs /eqP TSKs.
move: (INs) ⇒ NEQ.
eapply in_arrivals_implies_arrived_between in NEQ ⇒ //.
move: NEQ ⇒ /andP [_ JAs].
move: (BUSY) ⇒ [[ _ [QT [_ /andP [JAj _]]] _]].
apply QT ⇒ //; first exact: in_arrivals_implies_arrived.
apply H_policy_respects_sequential_tasks; first by rewrite TSK TSKs.
by apply leq_trans with t1; [lia |].
Qed.
interference_and_workload_consistent_with_sequential_tasks arr_seq sched tsk.
Proof.
move ⇒ j t1 t2 ARR /eqP TSK POS BUSY.
eapply instantiated_busy_interval_equivalent_busy_interval in BUSY ⇒ //.
eapply all_jobs_have_completed_equiv_workload_eq_service ⇒ //.
move ⇒ s INs /eqP TSKs.
move: (INs) ⇒ NEQ.
eapply in_arrivals_implies_arrived_between in NEQ ⇒ //.
move: NEQ ⇒ /andP [_ JAs].
move: (BUSY) ⇒ [[ _ [QT [_ /andP [JAj _]]] _]].
apply QT ⇒ //; first exact: in_arrivals_implies_arrived.
apply H_policy_respects_sequential_tasks; first by rewrite TSK TSKs.
by apply leq_trans with t1; [lia |].
Qed.
Since interfering and interfering workload are sufficient to define the busy window,
next, we reason about the bound on the length of the busy window.
Consider an arrival curve.
Consider a set of tasks that respects the arrival curve.
Variable ts : list Task.
Hypothesis H_taskset_respects_max_arrivals : taskset_respects_max_arrivals arr_seq ts.
Hypothesis H_taskset_respects_max_arrivals : taskset_respects_max_arrivals arr_seq ts.
Assume that all jobs come from this task set.
Consider a constant L such that...
... L is greater than 0, and...
L is the fixed point of the following equation.
Assume all jobs have a valid job cost.
Then, we prove that L is a bound on the length of the busy window.
Lemma instantiated_busy_intervals_are_bounded:
busy_intervals_are_bounded_by arr_seq sched tsk L.
Proof.
move ⇒ j ARR TSK POS.
edestruct busy_interval_from_total_workload_bound
with (Δ := L) as [t1 [t2 [T1 [T2 GGG]]]] ⇒ //.
{ by move ⇒ t; rewrite {2}H_fixed_point; apply total_workload_le_total_rbf. }
∃ t1, t2; split⇒ [//|]; split⇒ [//|].
by apply instantiated_busy_interval_equivalent_busy_interval.
Qed.
End BusyWindowBound.
End I_IW_correctness.
End JLFPInstantiation.
busy_intervals_are_bounded_by arr_seq sched tsk L.
Proof.
move ⇒ j ARR TSK POS.
edestruct busy_interval_from_total_workload_bound
with (Δ := L) as [t1 [t2 [T1 [T2 GGG]]]] ⇒ //.
{ by move ⇒ t; rewrite {2}H_fixed_point; apply total_workload_le_total_rbf. }
∃ t1, t2; split⇒ [//|]; split⇒ [//|].
by apply instantiated_busy_interval_equivalent_busy_interval.
Qed.
End BusyWindowBound.
End I_IW_correctness.
End JLFPInstantiation.
To preserve modularity and hide the implementation details of a
technical definition presented in this file, we make the
definition opaque. This way, we ensure that the system will treat
each of these definitions as a single entity.
Global Opaque another_hep_job_interference
another_task_hep_job_interference
ideal_jlfp_interference
ideal_jlfp_interfering_workload
cumulative_another_hep_job_interference
cumulative_another_task_hep_job_interference
cumulative_other_hep_jobs_interfering_workload
other_hep_jobs_interfering_workload.
another_task_hep_job_interference
ideal_jlfp_interference
ideal_jlfp_interfering_workload
cumulative_another_hep_job_interference
cumulative_another_task_hep_job_interference
cumulative_other_hep_jobs_interfering_workload
other_hep_jobs_interfering_workload.
We add some facts into the "Hint Database" basic_rt_facts, so Coq will be
able to apply them automatically where needed.
Global Hint Resolve
abstract_busy_interval_classic_quiet_time
abstract_busy_interval_classic_busy_interval_prefix
: basic_rt_facts.
abstract_busy_interval_classic_quiet_time
abstract_busy_interval_classic_busy_interval_prefix
: basic_rt_facts.