Built with
Alectryon , running Coq+SerAPI v8.20.0+0.20.0. Bubbles (
) indicate interactive fragments: hover for details, tap to reveal contents. Use
Ctrl+↑ Ctrl+↓ to navigate,
Ctrl+🖱️ to focus. On Mac, use
⌘ instead of
Ctrl .
Require Export prosa.analysis.facts.interference.[Loading ML file ssrmatching_plugin.cmxs (using legacy method) ... done ] [Loading ML file ssreflect_plugin.cmxs (using legacy method) ... done ] [Loading ML file ring_plugin.cmxs (using legacy method) ... done ] Serlib plugin: coq-elpi.elpi is not available: serlib support is missing.
Incremental checking for commands in this plugin will be impacted. [Loading ML file coq-elpi.elpi ... done ] [Loading ML file zify_plugin.cmxs (using legacy method) ... done ] [Loading ML file micromega_core_plugin.cmxs (using legacy method) ... done ] [Loading ML file micromega_plugin.cmxs (using legacy method) ... done ] [Loading ML file btauto_plugin.cmxs (using legacy method) ... done ] Notation "_ + _" was already used in scope nat_scope.
[notation-overridden,parsing,default]Notation "_ - _" was already used in scope nat_scope.
[notation-overridden,parsing,default]Notation "_ <= _" was already used in scope nat_scope.
[notation-overridden,parsing,default]Notation "_ < _" was already used in scope nat_scope.
[notation-overridden,parsing,default]Notation "_ >= _" was already used in scope nat_scope.
[notation-overridden,parsing,default]Notation "_ > _" was already used in scope nat_scope.
[notation-overridden,parsing,default]Notation "_ <= _ <= _" was already used in scope
nat_scope. [notation-overridden,parsing,default]Notation "_ < _ <= _" was already used in scope
nat_scope. [notation-overridden,parsing,default]Notation "_ <= _ < _" was already used in scope
nat_scope. [notation-overridden,parsing,default]Notation "_ < _ < _" was already used in scope
nat_scope. [notation-overridden,parsing,default]Notation "_ * _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Require Export prosa.analysis.abstract .IBF.supply_task.
Require Export prosa.analysis.facts.busy_interval.service_inversion.
(** * JLFP Instantiation of Interference and Interfering Workload for Restricted-Supply Uniprocessor *)
(** In this module we instantiate functions Interference and
Interfering Workload for the restricted-supply 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. *)
Section JLFPInstantiation .
(** Consider any type of tasks ... *)
Context {Task : TaskType}.
Context `{TaskCost Task}.
(** ... and any type of jobs associated with these tasks. *)
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
(** Consider any kind of fully supply-consuming unit-supply
uniprocessor model. *)
Context `{PState : ProcessorState Job}.
Hypothesis H_uniprocessor_proc_model : uniprocessor_model PState.
Hypothesis H_unit_supply_proc_model : unit_supply_proc_model PState.
Hypothesis H_consumed_supply_proc_model : fully_consuming_proc_model PState.
(** Consider any valid arrival sequence with consistent arrivals... *)
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_arrival_sequence : valid_arrival_sequence arr_seq.
(** ... and any valid uni-processor schedule of this arrival sequence... *)
Variable sched : schedule PState.
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.
(** 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.
(** Let [tsk] be any task. *)
Variable tsk : Task.
(** * Interference and Interfering Workload *)
(** In the following, we introduce definitions of interference and
interfering workload. *)
(** ** Instantiation of Interference *)
(** We say that job [j] incurs interference at time [t] iff it
cannot execute due to (1) the lack of supply at time [t], (2)
due to service inversion (i.e., a lower-priority job receiving
service at [t]), or higher-or-equal-priority job receiving
service. *)
#[local] Instance rs_jlfp_interference : Interference Job :=
{
interference (j : Job) (t : instant) :=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
}.
(** ** Instantiation of Interfering Workload *)
(** The interfering workload, in turn, is defined as the sum of the
blackout predicate, service inversion predicate, and interfering
workload of jobs with higher or equal priority. *)
#[local] Instance rs_jlfp_interfering_workload : InterferingWorkload Job :=
{
interfering_workload (j : Job) (t : instant) :=
is_blackout sched t
+ service_inversion arr_seq sched j t
+ other_hep_jobs_interfering_workload arr_seq j t
}.
(** ** 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.
As it was mentioned previously, instantiated functions of
interference and interfering workload usually do not have any
useful lemmas about them. However, it is possible to prove their
equivalence to the more conventional notions like service or
workload. Next we prove the equivalence between the
instantiations and conventional notions. *)
Section Equivalences .
(** We prove that we can split cumulative interference into three
parts: (1) blackout time, (2) cumulative service inversion,
and (3) cumulative interference from jobs with higher or equal
priority. *)
Lemma cumulative_interference_split :
forall j t1 t2 ,
cumulative_interference j t1 t2
= blackout_during sched t1 t2
+ cumulative_service_inversion arr_seq sched j t1 t2
+ cumulative_another_hep_job_interference arr_seq sched j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
cumulative_interference j t1 t2 =
blackout_during sched t1 t2 +
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_hep_job_interference arr_seq sched
j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
cumulative_interference j t1 t2 =
blackout_during sched t1 t2 +
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_hep_job_interference arr_seq sched
j t1 t2
rewrite /cumulative_interference /cumul_cond_interference /rs_jlfp_interference /cond_interference /interference.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
\sum_(t1 <= t < t2)
true &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) =
blackout_during sched t1 t2 +
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_hep_job_interference arr_seq sched
j t1 t2
move => j t1 t2; rewrite -big_split //= -big_split //=.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat
\sum_(t1 <= t < t2)
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) =
\sum_(t1 <= i < t2)
(is_blackout sched i +
service_inversion arr_seq sched j i +
another_hep_job_interference arr_seq sched j i)
apply /eqP; rewrite eqn_leq; apply /andP; split ; rewrite leq_sum//; move => t _.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t <=
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t <=
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
by case is_blackout, service_inversion, another_hep_job_interference. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
have [BL|SUP] := blackout_or_supply sched t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat BL : is_blackout sched t
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat BL : is_blackout sched t
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
rewrite BL //= -addnA addnC addn1 ltnS leqn0 addn_eq0.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat BL : is_blackout sched t
(service_inversion arr_seq sched j t == 0 ) &&
(another_hep_job_interference arr_seq sched j t == 0 )
by apply /andP; split ;
[rewrite eqb0 blackout_implies_no_service_inversion
| rewrite eqb0 no_hep_job_interference_without_supply]. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t
is_blackout sched t +
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
rewrite /is_blackout SUP //= add0n.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
have [IDLE|[s SCHEDs]] := scheduled_at_cases arr_seq ltac :(eauto ) sched ltac :(eauto ) ltac :(eauto ) t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t IDLE : is_idle arr_seq sched t
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t IDLE : is_idle arr_seq sched t
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
by rewrite -[service_inversion _ _ _ _]negbK idle_implies_no_service_inversion //=. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t
rewrite (service_inversion_supply_sched _ _ _ _ _ _ _ _ _ _ s) //
(interference_ahep_def _ _ _ _ _ _ _ _ _ _ s) //
/another_hep_job.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
~~ hep_job_at t s j + hep_job s j && (s != j) <=
~~ hep_job_at t s j || hep_job s j && (s != j)
have [EQ|NEQ] := eqVneq s j.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t EQ : s = j
~~ hep_job_at t s j + hep_job s j && ~~ true <=
~~ hep_job_at t s j || hep_job s j && ~~ true
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t EQ : s = j
~~ hep_job_at t s j + hep_job s j && ~~ true <=
~~ hep_job_at t s j || hep_job s j && ~~ true
by rewrite andbF orbF addn0. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQ : s != j
~~ hep_job_at t s j + hep_job s j && ~~ false <=
~~ hep_job_at t s j || hep_job s j && ~~ false
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2, t : nat SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQ : s != j
~~ hep_job_at t s j + hep_job s j && ~~ false <=
~~ hep_job_at t s j || hep_job s j && ~~ false
by unfold hep_job_at, JLFP_to_JLDP, hep_job; rewrite andbT; case (JLFP s j) => //. }
}
}
}
Qed .
(** Similarly, we prove that we can split cumulative interfering
workload into three parts: (1) blackout time, (2) cumulative
service inversion, and (3) cumulative interfering workload
from jobs with higher or equal priority. *)
Lemma cumulative_interfering_workload_split :
forall j t1 t2 ,
cumulative_interfering_workload j t1 t2 =
blackout_during sched t1 t2
+ cumulative_service_inversion arr_seq sched j t1 t2
+ cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
cumulative_interfering_workload j t1 t2 =
blackout_during sched t1 t2 +
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_other_hep_jobs_interfering_workload arr_seq
j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
cumulative_interfering_workload j t1 t2 =
blackout_during sched t1 t2 +
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_other_hep_jobs_interfering_workload arr_seq
j t1 t2
by move => j t1 t2; rewrite -big_split //= -big_split //=.
Qed .
(** Let <<[t1, t2)>> be a time interval and let [j] be any job of
task [tsk] that is not completed by time [t2]. Then cumulative
interference received due jobs of other tasks executing can be
bounded by the sum of the cumulative service inversion of job
[j] and the cumulative interference incurred by task [tsk] due
to other tasks. *)
Lemma cumulative_task_interference_split :
forall j t1 t2 ,
arrives_in arr_seq j ->
job_of_task tsk j ->
~~ completed_by sched j t2 ->
cumul_cond_interference (nonself_intra arr_seq sched) j t1 t2
<= cumulative_service_inversion arr_seq sched j t1 t2
+ cumulative_another_task_hep_job_interference arr_seq sched j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 : nat) (t2 : instant),
arrives_in arr_seq j ->
job_of_task tsk j ->
~~ completed_by sched j t2 ->
cumul_cond_interference (nonself_intra arr_seq sched)
j t1 t2 <=
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_task_hep_job_interference arr_seq
sched j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 : nat) (t2 : instant),
arrives_in arr_seq j ->
job_of_task tsk j ->
~~ completed_by sched j t2 ->
cumul_cond_interference (nonself_intra arr_seq sched)
j t1 t2 <=
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_task_hep_job_interference arr_seq
sched j t1 t2
move => j t1 R ARR TSK NCOMPL.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R
cumul_cond_interference (nonself_intra arr_seq sched)
j t1 R <=
cumulative_service_inversion arr_seq sched j t1 R +
cumulative_another_task_hep_job_interference arr_seq
sched j t1 R
rewrite /cumul_task_interference /cumul_cond_interference.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R
\sum_(t1 <= t < R)
cond_interference (nonself_intra arr_seq sched) j t <=
cumulative_service_inversion arr_seq sched j t1 R +
cumulative_another_task_hep_job_interference arr_seq
sched j t1 R
rewrite -big_split //= big_seq_cond [leqRHS]big_seq_cond.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R
\sum_(i <- index_iota t1 R | (i \in index_iota t1 R) &&
true)
cond_interference (nonself_intra arr_seq sched) j i <=
\sum_(i <- index_iota t1 R | (i \in index_iota t1 R) &&
true)
(service_inversion arr_seq sched j i +
another_task_hep_job_interference arr_seq sched j
i)
apply leq_sum; move => t /andP [IN _].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R
cond_interference (nonself_intra arr_seq sched) j t <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
rewrite /cond_interference /nonself /interference /rs_jlfp_interference /nonself_intra.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R
nonself arr_seq sched j t && has_supply sched t &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
have [SUP|NOSUP] := boolP (has_supply sched t); last first .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R NOSUP : ~~ has_supply sched t
nonself arr_seq sched j t && false &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R NOSUP : ~~ has_supply sched t
nonself arr_seq sched j t && false &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
by move : (NOSUP); rewrite /is_blackout => -> //=; rewrite andbT andbF //. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t
nonself arr_seq sched j t && true &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t
nonself arr_seq sched j t && true &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
move : (SUP); rewrite /is_blackout => ->; rewrite andbT //=.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t
nonself arr_seq sched j t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
have [IDLE|[s SCHEDs]] := scheduled_at_cases arr_seq ltac :(eauto ) sched ltac :(eauto ) ltac :(eauto ) t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t IDLE : is_idle arr_seq sched t
nonself arr_seq sched j t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t IDLE : is_idle arr_seq sched t
nonself arr_seq sched j t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
rewrite /nonself.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t IDLE : is_idle arr_seq sched t
~~ task_served_at arr_seq sched (job_task j) t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
by rewrite -[service_inversion _ _ _ _]negbK
idle_implies_no_service_inversion //;
rewrite_neg no_hep_job_interference_when_idle;
rewrite_neg no_hep_task_interference_when_idle; rewrite andbF. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
nonself arr_seq sched j t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
nonself arr_seq sched j t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
rewrite /nonself; move : (TSK) => /eqP ->.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
~~ task_served_at arr_seq sched tsk t &&
(service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_task_hep_job_interference arr_seq sched j t
erewrite task_served_at_eq_job_of_task => //; erewrite service_inversion_supply_sched => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
~~ job_of_task tsk s &&
(~~ hep_job_at t s j
|| another_hep_job_interference arr_seq sched j t) <=
~~ hep_job_at t s j +
another_task_hep_job_interference arr_seq sched j t
erewrite interference_ahep_def => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || another_hep_job s j) <=
~~ hep_job_at t s j +
another_task_hep_job_interference arr_seq sched j t
erewrite interference_athep_def => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || another_hep_job s j) <=
~~ hep_job_at t s j + another_task_hep_job s j
rewrite /another_hep_job /another_task_hep_job.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && (s != j)) <=
~~ hep_job_at t s j +
hep_job s j && (job_task s != job_task j)
have [EQj|NEQj] := eqVneq s j.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t EQj : s = j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ true) <=
~~ hep_job_at t s j +
hep_job s j && (job_task s != job_task j)
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t EQj : s = j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ true) <=
~~ hep_job_at t s j +
hep_job s j && (job_task s != job_task j)
by subst ; rewrite /job_of_task; move : TSK => /eqP => ->; rewrite H_priority_is_reflexive eq_refl. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ false) <=
~~ hep_job_at t s j +
hep_job s j && (job_task s != job_task j)
have [/eqP EQt|NEQt] := eqVneq (job_task s) (job_task j).Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j EQt : job_task s == job_task j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ false) <=
~~ hep_job_at t s j + hep_job s j && ~~ true
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j EQt : job_task s == job_task j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ false) <=
~~ hep_job_at t s j + hep_job s j && ~~ true
rewrite /job_of_task; move : TSK => /eqP <-; rewrite EQt.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j EQt : job_task s == job_task j
~~ true &&
(~~ hep_job_at t s j || hep_job s j && ~~ false) <=
~~ hep_job_at t s j + hep_job s j && ~~ true
by apply /eqP; rewrite andbF andFb addn0 //=. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j NEQt : job_task s != job_task j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ false) <=
~~ hep_job_at t s j + hep_job s j && ~~ false
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j NEQt : job_task s != job_task j
~~ job_of_task tsk s &&
(~~ hep_job_at t s j || hep_job s j && ~~ false) <=
~~ hep_job_at t s j + hep_job s j && ~~ false
unfold hep_job_at, JLFP_to_JLDP, hep_job.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1 : nat R : instant ARR : arrives_in arr_seq j TSK : job_of_task tsk j NCOMPL : ~~ completed_by sched j R t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 R SUP : has_supply sched t s : Job SCHEDs : scheduled_at sched s t NEQj : s != j NEQt : job_task s != job_task j
~~ job_of_task tsk s &&
(~~ JLFP s j || JLFP s j && ~~ false) <=
~~ JLFP s j + JLFP s j && ~~ false
by rewrite /job_of_task; move : TSK => /eqP <-; rewrite NEQt //= andbT; case : JLFP. }
}
}
Qed .
(** We also show that the cumulative intra-supply interference can
be split into the sum of the cumulative service inversion and
cumulative interference incurred by the job due to other
higher-or-equal priority jobs. *)
Lemma cumulative_intra_interference_split :
forall j t1 t2 ,
cumul_cond_interference (fun (_j : Job) (t : instant) => has_supply sched t) j t1 t2
<= cumulative_service_inversion arr_seq sched j t1 t2
+ cumulative_another_hep_job_interference arr_seq sched j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
cumul_cond_interference (fun => [eta has_supply sched])
j t1 t2 <=
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_hep_job_interference arr_seq sched
j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task
forall (j : Job) (t1 t2 : nat),
cumul_cond_interference (fun => [eta has_supply sched])
j t1 t2 <=
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_hep_job_interference arr_seq sched
j t1 t2
move => j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat
cumul_cond_interference (fun => [eta has_supply sched])
j t1 t2 <=
cumulative_service_inversion arr_seq sched j t1 t2 +
cumulative_another_hep_job_interference arr_seq sched
j t1 t2
rewrite /cumul_cond_interference -big_split //= big_seq_cond [leqRHS]big_seq_cond.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat
\sum_(i <- index_iota t1 t2 | (i \in index_iota t1 t2) &&
true)
cond_interference (fun => [eta has_supply sched]) j
i <=
\sum_(i <- index_iota t1 t2 | (i \in index_iota t1 t2) &&
true)
(service_inversion arr_seq sched j i +
another_hep_job_interference arr_seq sched j i)
apply leq_sum; move => t /andP [IN _].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2
cond_interference (fun => [eta has_supply sched]) j t <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
rewrite /cond_interference /nonself /interference /rs_jlfp_interference.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2
has_supply sched t &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
have [SUP|NOSUP] := boolP (has_supply sched t); last first .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2 NOSUP : ~~ has_supply sched t
false &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2 NOSUP : ~~ has_supply sched t
false &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
by move : (NOSUP); rewrite /is_blackout => -> //=; rewrite andbT andbF //. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2 SUP : has_supply sched t
true &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2 SUP : has_supply sched t
true &&
(is_blackout sched t
|| service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t) <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
move : (SUP); rewrite /is_blackout => ->; rewrite andTb //=.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2 SUP : has_supply sched t
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
have L : forall (a b : bool), a || b <= a + b by clear => [] [] [].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task j : Job t1, t2 : nat t : Datatypes_nat__canonical__eqtype_Equality IN : t \in index_iota t1 t2 SUP : has_supply sched t L : forall a b : bool, a || b <= a + b
service_inversion arr_seq sched j t
|| another_hep_job_interference arr_seq sched j t <=
service_inversion arr_seq sched j t +
another_hep_job_interference arr_seq sched j t
by apply L. }
Qed .
(** In this section, we prove that the (abstract) cumulative
interfering workload due to other higher-or-equal priority
jobs is equal to the conventional workload (from other
higher-or-equal priority jobs). *)
Section InstantiatedWorkloadEquivalence .
(** Let <<[t1, t2)>> be any time interval. *)
Variables t1 t2 : instant.
(** Consider any job [j] of [tsk]. *)
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
(** The cumulative interfering workload (w.r.t. [j]) due to
other higher-or-equal priority jobs is equal to the
conventional workload from other higher-or-equal priority
jobs. *)
Lemma cumulative_iw_hep_eq_workload_of_ohep :
cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2
= workload_of_other_hep_jobs arr_seq j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j
cumulative_other_hep_jobs_interfering_workload arr_seq
j t1 t2 = workload_of_other_hep_jobs arr_seq j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j
cumulative_other_hep_jobs_interfering_workload arr_seq
j t1 t2 = workload_of_other_hep_jobs arr_seq j t1 t2
rewrite /cumulative_other_hep_jobs_interfering_workload /workload_of_other_hep_jobs.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j
\sum_(t1 <= t < t2)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 t2)
case NEQ: (t1 < t2); last first .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j NEQ : (t1 < t2) = false
\sum_(t1 <= t < t2)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 t2)
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j NEQ : (t1 < t2) = false
\sum_(t1 <= t < t2)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 t2)
move : NEQ => /negP /negP; rewrite -leqNgt; move => NEQ.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j NEQ : t2 <= t1
\sum_(t1 <= t < t2)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 t2)
rewrite big_geq // /arrivals_between big_geq //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j NEQ : t2 <= t1
0 = workload_of_jobs (another_hep_job^~ j) [::]
by rewrite /workload_of_jobs big_nil. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j NEQ : (t1 < t2) = true
\sum_(t1 <= t < t2)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 t2)
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1, t2 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j NEQ : (t1 < t2) = true
\sum_(t1 <= t < t2)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 t2)
interval_to_duration t1 t2 k. Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j k : nat
\sum_(t1 <= t < t1 + k)
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k))
elim : k => [|k IHk].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j
\sum_(t1 <= t < t1 + 0 )
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + 0 ))
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j
\sum_(t1 <= t < t1 + 0 )
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + 0 ))
rewrite !addn0 big_geq// /arrivals_between big_geq//.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j
0 = workload_of_jobs (another_hep_job^~ j) [::]
by rewrite /workload_of_jobs big_nil.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j k : nat IHk : \sum_(t1 <= t <
t1 + k)
other_hep_jobs_interfering_workload arr_seq j
t =
workload_of_jobs
(another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k))
\sum_(t1 <= t < t1 + k.+1 )
other_hep_jobs_interfering_workload arr_seq j t =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k.+1 ))
rewrite addnS big_nat_recr //=; last by rewrite leq_addr.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j k : nat IHk : \sum_(t1 <= t <
t1 + k)
other_hep_jobs_interfering_workload arr_seq j
t =
workload_of_jobs
(another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k))
\sum_(t1 <= i < t1 + k)
other_hep_jobs_interfering_workload arr_seq j i +
other_hep_jobs_interfering_workload arr_seq j (t1 + k) =
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k).+1 )
rewrite IHk /arrivals_between big_nat_recr //=.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j k : nat IHk : \sum_(t1 <= t <
t1 + k)
other_hep_jobs_interfering_workload arr_seq j
t =
workload_of_jobs
(another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k))
workload_of_jobs (another_hep_job^~ j)
(\cat_(t1<=t<t1 + k)arrivals_at arr_seq t) +
other_hep_jobs_interfering_workload arr_seq j (t1 + k) =
workload_of_jobs (another_hep_job^~ j)
(\cat_(t1<=i<t1 + k)arrivals_at arr_seq i ++
arrivals_at arr_seq (t1 + k))
+ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j k : nat IHk : \sum_(t1 <= t <
t1 + k)
other_hep_jobs_interfering_workload arr_seq j
t =
workload_of_jobs
(another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k))
workload_of_jobs (another_hep_job^~ j)
(\cat_(t1<=t<t1 + k)arrivals_at arr_seq t) +
other_hep_jobs_interfering_workload arr_seq j (t1 + k) =
workload_of_jobs (another_hep_job^~ j)
(\cat_(t1<=i<t1 + k)arrivals_at arr_seq i ++
arrivals_at arr_seq (t1 + k))
by rewrite /workload_of_jobs big_cat.
+ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task t1 : instant j : Job H_j_arrives : arrives_in arr_seq j H_job_of_tsk : job_of_task tsk j k : nat IHk : \sum_(t1 <= t <
t1 + k)
other_hep_jobs_interfering_workload arr_seq j
t =
workload_of_jobs
(another_hep_job^~ j)
(arrivals_between arr_seq t1 (t1 + k))
t1 <= t1 + k
by rewrite leq_addr. }
Qed .
End InstantiatedWorkloadEquivalence .
(** 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. *)
Section BusyIntervalEquivalence .
(** 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 := classical.quiet_time arr_seq sched.
Let quiet_time_ab := abstract .definitions.quiet_time sched.
(** Same for the two notions of a busy interval prefix ... *)
Let busy_interval_prefix_cl := classical.busy_interval_prefix arr_seq sched.
Let busy_interval_prefix_ab := abstract .definitions.busy_interval_prefix sched.
(** ... and the two notions of a busy interval. *)
Let busy_interval_cl := classical.busy_interval arr_seq sched.
Let busy_interval_ab := abstract .definitions.busy_interval sched.
(** Consider any job [j] of [tsk]. *)
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq 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 :
forall t , quiet_time_cl j t -> quiet_time_ab j t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t : instant,
quiet_time_cl j t -> quiet_time_ab j t
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t : instant,
quiet_time_cl j t -> quiet_time_ab j t
have zero_is_quiet_time: forall j , quiet_time_cl j 0 .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall j : Job, quiet_time_cl j 0
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall j : Job, quiet_time_cl j 0
by move => jhp ARR HP AB; move : AB; rewrite /arrived_before ltn0. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
forall t : instant,
quiet_time_cl j t -> quiet_time_ab j t
move => t QT; apply /andP; split ; last first .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
~~ pending_earlier_and_at sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
~~ pending_earlier_and_at sched j t
rewrite negb_and negbK; apply /orP.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
~~ arrived_before j t \/ completed_by sched j t
by case ARR: (arrived_before j t); [right | left ]; [apply QT | ]. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t
erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
cumulative_another_hep_job_interference arr_seq sched
j 0 t ==
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t
rewrite cumulative_i_ohep_eq_service_of_ohep //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
service_of_other_hep_jobs arr_seq sched j 0 t ==
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t
rewrite //= cumulative_iw_hep_eq_workload_of_ohep eq_sym; apply /eqP.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
workload_of_other_hep_jobs arr_seq j 0 t =
service_of_other_hep_jobs arr_seq sched j 0 t
apply all_jobs_have_completed_equiv_workload_eq_service => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t
forall j0 : Job,
j0 \in arrivals_between arr_seq 0 t ->
another_hep_job j0 j -> completed_by sched j0 t
move => j0 IN HEP; apply QT.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t j0 : Job IN : j0 \in arrivals_between arr_seq 0 t HEP : another_hep_job j0 j
arrives_in arr_seq j0
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t j0 : Job IN : j0 \in arrivals_between arr_seq 0 t HEP : another_hep_job j0 j
arrives_in arr_seq j0
by apply in_arrivals_implies_arrived in IN.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t j0 : Job IN : j0 \in arrivals_between arr_seq 0 t HEP : another_hep_job j0 j
hep_job j0 j
by move : HEP => /andP [HEP HP].
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant QT : quiet_time_cl j t j0 : Job IN : j0 \in arrivals_between arr_seq 0 t HEP : another_hep_job j0 j
arrived_before j0 t
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 :
forall t , quiet_time_ab j t -> quiet_time_cl j t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t : instant,
quiet_time_ab j t -> quiet_time_cl j t
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t : instant,
quiet_time_ab j t -> quiet_time_cl j t
have zero_is_quiet_time: forall j , quiet_time_cl j 0 .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall j : Job, quiet_time_cl j 0
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall j : Job, quiet_time_cl j 0
by move => jhp ARR HP AB; move : AB; rewrite /arrived_before ltn0. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
forall t : instant,
quiet_time_ab j t -> quiet_time_cl j t
move => t /andP [T0 T1] jhp ARR HP ARB.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
completed_by sched jhp t
eapply all_jobs_have_completed_equiv_workload_eq_service with
(P := fun jhp => hep_job jhp j) (t1 := 0 ) (t2 := t) => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
workload_of_jobs (hep_job^~ j)
(arrivals_between arr_seq 0 t) =
service_of_jobs sched (hep_job^~ j)
(arrivals_between arr_seq 0 t) 0 t
erewrite service_of_jobs_case_on_pred with (P2 := fun j' => j' != j).Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
workload_of_jobs (hep_job^~ j)
(arrivals_between arr_seq 0 t) =
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && (j0 != j))
(arrivals_between arr_seq 0 t) 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
erewrite workload_of_jobs_case_on_pred with (P' := fun j' => j' != j) => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
workload_of_jobs
(fun j0 : Job => hep_job j0 j && (j0 != j))
(arrivals_between arr_seq 0 t) +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && (j0 != j))
(arrivals_between arr_seq 0 t) 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
replace ((fun j0 : Job => hep_job j0 j && (j0 != j))) with (another_hep_job^~j); last by rewrite /another_hep_job.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq 0 t) +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
service_of_jobs sched (another_hep_job^~ j)
(arrivals_between arr_seq 0 t) 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
rewrite -/(service_of_other_hep_jobs arr_seq sched j 0 t) -cumulative_i_ohep_eq_service_of_ohep //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
workload_of_jobs (another_hep_job^~ j)
(arrivals_between arr_seq 0 t) +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
rewrite -/(workload_of_other_hep_jobs arr_seq j 0 t) -cumulative_iw_hep_eq_workload_of_ohep //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t T1 : ~~ pending_earlier_and_at sched j t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
move : T1; rewrite negb_and => /orP [NA | /negPn COMP].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
have PRED__eq: {in arrivals_between arr_seq 0 t, (fun j__copy : Job => hep_job j__copy j && ~~ (j__copy != j)) =1 pred0}.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t
{in arrivals_between arr_seq 0 t,
(fun j__copy : Job =>
hep_job j__copy j && ~~ (j__copy != j)) =1 pred0}
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t
{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.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t j__copy : Job IN : j__copy \in arrivals_between arr_seq 0 t
~~ (hep_job j__copy j && ~~ (j__copy != j))
rewrite negb_and; apply /orP; right ; apply /negPn/eqP => EQ; subst j__copy.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t IN : j \in arrivals_between arr_seq 0 t
False
move : NA => /negP CONTR; apply : CONTR.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t IN : j \in arrivals_between arr_seq 0 t
arrived_before j t
by apply in_arrivals_implies_arrived_between in IN. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j__copy : Job =>
hep_job j__copy j && ~~ (j__copy != j)) =1
pred0}
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
erewrite service_of_jobs_equiv_pred with (P2 := pred0) => [|//].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j__copy : Job =>
hep_job j__copy j && ~~ (j__copy != j)) =1
pred0}
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched pred0
(arrivals_between arr_seq 0 t) 0 t
erewrite workload_of_jobs_equiv_pred with (P' := pred0) => [|//].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j__copy : Job =>
hep_job j__copy j && ~~ (j__copy != j)) =1
pred0}
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs pred0 (arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched pred0
(arrivals_between arr_seq 0 t) 0 t
move : T0; erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l => /eqP EQ.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j__copy : Job =>
hep_job j__copy j && ~~ (j__copy != j)) =1
pred0} EQ : cumulative_another_hep_job_interference arr_seq
sched j 0 t =
cumulative_other_hep_jobs_interfering_workload
arr_seq j 0 t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs pred0 (arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched pred0
(arrivals_between arr_seq 0 t) 0 t
rewrite EQ; clear EQ; apply /eqP; rewrite eqn_add2l.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t NA : ~~ arrived_before j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j__copy : Job =>
hep_job j__copy j && ~~ (j__copy != j)) =1
pred0}
workload_of_jobs pred0 (arrivals_between arr_seq 0 t) ==
service_of_jobs sched pred0
(arrivals_between arr_seq 0 t) 0 t
by erewrite workload_of_jobs_pred0, service_of_jobs_pred0.
} Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
have PRED__eq: {in arrivals_between arr_seq 0 t, (fun j0 : Job => hep_job j0 j && ~~ (j0 != j)) =1 eq_op j}.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t
{in arrivals_between arr_seq 0 t,
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t
{in arrivals_between arr_seq 0 t,
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
move => j__copy IN.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t j__copy : Job IN : j__copy \in arrivals_between arr_seq 0 t
hep_job j__copy j && ~~ (j__copy != j) =
(j == j__copy)
replace (~~ (j__copy != j)) with (j__copy == j); last by case : (j__copy == j).Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t j__copy : Job IN : j__copy \in arrivals_between arr_seq 0 t
hep_job j__copy j && (j__copy == j) = (j == j__copy)
rewrite eq_sym; destruct (j == j__copy) eqn :EQ; last by rewrite Bool.andb_false_r.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t j__copy : Job IN : j__copy \in arrivals_between arr_seq 0 t EQ : (j == j__copy) = true
hep_job j__copy j && true = true
by move : EQ => /eqP EQ; rewrite Bool.andb_true_r; apply /eqP; rewrite eqb_id; subst . } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) 0 t
erewrite service_of_jobs_equiv_pred with (P2 := eq_op j) => [|//].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs
(fun j0 : Job => hep_job j0 j && ~~ (j0 != j))
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched (eq_op j)
(arrivals_between arr_seq 0 t) 0 t
erewrite workload_of_jobs_equiv_pred with (P' := eq_op j) => [|//].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant T0 : cumulative_interference j 0 t ==
cumulative_interfering_workload j 0 t jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs (eq_op j)
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched (eq_op j)
(arrivals_between arr_seq 0 t) 0 t
move : T0; erewrite cumulative_interference_split, cumulative_interfering_workload_split; rewrite eqn_add2l => /eqP EQ.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j} EQ : cumulative_another_hep_job_interference arr_seq
sched j 0 t =
cumulative_other_hep_jobs_interfering_workload
arr_seq j 0 t
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t +
workload_of_jobs (eq_op j)
(arrivals_between arr_seq 0 t) =
cumulative_another_hep_job_interference arr_seq sched
j 0 t +
service_of_jobs sched (eq_op j)
(arrivals_between arr_seq 0 t) 0 t
rewrite EQ; clear EQ; apply /eqP; rewrite eqn_add2l.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
workload_of_jobs (eq_op j)
(arrivals_between arr_seq 0 t) ==
service_of_jobs sched (eq_op j)
(arrivals_between arr_seq 0 t) 0 t
apply /eqP; eapply all_jobs_have_completed_equiv_workload_eq_service with
(P := eq_op j) (t1 := 0 ) (t2 := t) => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
forall j0 : Job,
j0 \in arrivals_between arr_seq 0 t ->
j == j0 -> completed_by sched j0 t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j zero_is_quiet_time : forall j : Job, quiet_time_cl j 0 t : instant jhp : Job ARR : arrives_in arr_seq jhp HP : hep_job jhp j ARB : arrived_before jhp t COMP : completed_by sched j t PRED__eq : {in arrivals_between arr_seq 0 t,
(fun j0 : Job =>
hep_job j0 j && ~~ (j0 != j)) =1
eq_op j}
forall j0 : Job,
j0 \in arrivals_between arr_seq 0 t ->
j == j0 -> completed_by sched j0 t
by move => j__copy _ /eqP EQ; subst j__copy. }
}
Qed .
(** The equivalence trivially follows from the lemmas above. *)
Corollary instantiated_quiet_time_equivalent_quiet_time :
forall t ,
quiet_time_cl j t <-> quiet_time_ab j t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t : instant,
quiet_time_cl j t <-> quiet_time_ab j t
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t : instant,
quiet_time_cl j t <-> quiet_time_ab j t
move => ?; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j _t_ : instant
quiet_time_cl j _t_ -> quiet_time_ab j _t_
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j _t_ : instant
quiet_time_cl j _t_ -> quiet_time_ab j _t_
by apply quiet_time_cl_implies_quiet_time_ab.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j _t_ : instant
quiet_time_ab j _t_ -> quiet_time_cl j _t_
by apply quiet_time_ab_implies_quiet_time_cl.
Qed .
(** Based on that, we prove that the concept of a 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 :
forall t1 t2 , busy_interval_prefix_cl j t1 t2 <-> busy_interval_prefix_ab j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_prefix_cl j t1 t2 <->
busy_interval_prefix_ab j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_prefix_cl j t1 t2 <->
busy_interval_prefix_ab j t1 t2
move => t1 t2; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_prefix_cl j t1 t2 ->
busy_interval_prefix_ab j t1 t2
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_prefix_cl j t1 t2 ->
busy_interval_prefix_ab j t1 t2
move => [NEQ [QTt1 [NQT REL]]].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ : t1 < t2 QTt1 : quiet_time arr_seq sched j t1 NQT : forall t : nat,
t1 < t < t2 -> ~ quiet_time arr_seq sched j tREL : t1 <= job_arrival j < t2
busy_interval_prefix_ab j t1 t2
split => [//|]; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ : t1 < t2 QTt1 : quiet_time arr_seq sched j t1 NQT : forall t : nat,
t1 < t < t2 -> ~ quiet_time arr_seq sched j tREL : t1 <= job_arrival j < t2
definitions.quiet_time sched j t1
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ : t1 < t2 QTt1 : quiet_time arr_seq sched j t1 NQT : forall t : nat,
t1 < t < t2 -> ~ quiet_time arr_seq sched j tREL : t1 <= job_arrival j < t2
definitions.quiet_time sched j t1
by apply instantiated_quiet_time_equivalent_quiet_time in QTt1.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ : t1 < t2 QTt1 : quiet_time arr_seq sched j t1 NQT : forall t : nat,
t1 < t < t2 -> ~ quiet_time arr_seq sched j tREL : t1 <= job_arrival j < t2
forall t : nat,
t1 < t < t2 -> ~ definitions.quiet_time sched j t
by move => t NE QT; eapply NQT; eauto 2 ; apply instantiated_quiet_time_equivalent_quiet_time.
} Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_prefix_ab j t1 t2 ->
busy_interval_prefix_cl j t1 t2
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_prefix_ab j t1 t2 ->
busy_interval_prefix_cl j t1 t2
move => [/andP [NEQ1 NEQ2] [QTt1 NQT]].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ1 : t1 <= job_arrival j NEQ2 : job_arrival j < t2 QTt1 : definitions.quiet_time sched j t1 NQT : forall t : nat,
t1 < t < t2 ->
~ definitions.quiet_time sched j t
busy_interval_prefix_cl j t1 t2
split ; [ | split ; [ |split ] ].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ1 : t1 <= job_arrival j NEQ2 : job_arrival j < t2 QTt1 : definitions.quiet_time sched j t1 NQT : forall t : nat,
t1 < t < t2 ->
~ definitions.quiet_time sched j t
t1 < t2
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ1 : t1 <= job_arrival j NEQ2 : job_arrival j < t2 QTt1 : definitions.quiet_time sched j t1 NQT : forall t : nat,
t1 < t < t2 ->
~ definitions.quiet_time sched j t
t1 < t2
by apply leq_ltn_trans with (job_arrival j).
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ1 : t1 <= job_arrival j NEQ2 : job_arrival j < t2 QTt1 : definitions.quiet_time sched j t1 NQT : forall t : nat,
t1 < t < t2 ->
~ definitions.quiet_time sched j t
quiet_time arr_seq sched j t1
by eapply instantiated_quiet_time_equivalent_quiet_time.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ1 : t1 <= job_arrival j NEQ2 : job_arrival j < t2 QTt1 : definitions.quiet_time sched j t1 NQT : forall t : nat,
t1 < t < t2 ->
~ definitions.quiet_time sched j t
forall t : nat,
t1 < t < t2 -> ~ quiet_time arr_seq sched j t
by move => t NEQ QT; eapply NQT; eauto 2 ; eapply instantiated_quiet_time_equivalent_quiet_time in QT.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant NEQ1 : t1 <= job_arrival j NEQ2 : job_arrival j < t2 QTt1 : definitions.quiet_time sched j t1 NQT : forall t : nat,
t1 < t < t2 ->
~ definitions.quiet_time sched j t
t1 <= job_arrival j < t2
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 :
forall t1 t2 , busy_interval_cl j t1 t2 <-> busy_interval_ab j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_cl j t1 t2 <-> busy_interval_ab j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_cl j t1 t2 <-> busy_interval_ab j t1 t2
move => t1 t2; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_cl j t1 t2 -> busy_interval_ab j t1 t2
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_cl j t1 t2 -> busy_interval_ab j t1 t2
move => [PREF QTt2]; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant PREF : busy_interval_prefix arr_seq sched j t1 t2 QTt2 : quiet_time arr_seq sched j t2
definitions.busy_interval_prefix sched j t1 t2
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant PREF : busy_interval_prefix arr_seq sched j t1 t2 QTt2 : quiet_time arr_seq sched j t2
definitions.busy_interval_prefix sched j t1 t2
by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant PREF : busy_interval_prefix arr_seq sched j t1 t2 QTt2 : quiet_time arr_seq sched j t2
definitions.quiet_time sched j t2
by eapply instantiated_quiet_time_equivalent_quiet_time in QTt2.
} Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_ab j t1 t2 -> busy_interval_cl j t1 t2
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant
busy_interval_ab j t1 t2 -> busy_interval_cl j t1 t2
move => [PREF QTt2]; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant PREF : definitions.busy_interval_prefix sched j t1 t2 QTt2 : definitions.quiet_time sched j t2
busy_interval_prefix arr_seq sched j t1 t2
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant PREF : definitions.busy_interval_prefix sched j t1 t2 QTt2 : definitions.quiet_time sched j t2
busy_interval_prefix arr_seq sched j t1 t2
by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j t1, t2 : instant PREF : definitions.busy_interval_prefix sched j t1 t2 QTt2 : definitions.quiet_time sched j t2
quiet_time arr_seq sched j t2
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 :
forall t1 t2 ,
busy_interval_ab j t1 t2 -> quiet_time_cl j t1.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_ab j t1 t2 -> quiet_time_cl j t1
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_ab j t1 t2 -> quiet_time_cl j t1
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 :
forall t1 t2 ,
busy_interval_ab j t1 t2 -> busy_interval_prefix_cl j t1 t2.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_ab j t1 t2 ->
busy_interval_prefix_cl j t1 t2
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task quiet_time_cl := quiet_time arr_seq sched : Job -> instant -> Prop quiet_time_ab := definitions.quiet_time sched : Job -> instant -> bool busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_cl := busy_interval arr_seq sched : Job -> instant -> instant -> Prop busy_interval_ab := definitions.busy_interval sched : Job -> instant -> instant -> Prop j : Job H_j_arrives : arrives_in arr_seq j
forall t1 t2 : instant,
busy_interval_ab j t1 t2 ->
busy_interval_prefix_cl j t1 t2
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. *)
Section I_IW_correctness .
(** Consider work-bearing readiness. *)
Context `{!JobReady Job PState}.
Hypothesis H_work_bearing_readiness : work_bearing_readiness arr_seq sched.
(** Assume that the schedule is valid and work-conserving. *)
Hypothesis H_sched_valid : valid_schedule sched arr_seq.
(** 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.
(** ... as well as notions of busy interval prefix. *)
Let busy_interval_prefix_ab := definitions.busy_interval_prefix sched.
Let busy_interval_prefix_cl := classical.busy_interval_prefix arr_seq 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. *)
Hypothesis H_work_conserving : work_conserving_cl.
(** 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. *)
Section 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.
(** Let the busy interval of the job be <<[t1, t2)>>. *)
Variable t1 t2 : instant.
Hypothesis H_busy_interval_prefix : busy_interval_prefix_ab j t1 t2.
(** Consider a time [t] inside the busy interval of the job. *)
Variable t : instant.
Hypothesis H_t_in_busy_interval : t1 <= t < t2.
(** First, we note that, similarly to the ideal uni-processor
case, there is no idle time inside of a busy interval. That
is, there is a job scheduled at time [t]. *)
Local Lemma busy_implies_not_idle :
exists j , scheduled_at sched j t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2
exists j : Job, scheduled_at sched j t
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2
exists j : Job, scheduled_at sched j t
have [IDLE|[s SCHEDs]] := scheduled_at_cases arr_seq ltac :(eauto ) sched ltac :(eauto ) ltac :(eauto ) t; last by (exists s ).Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 IDLE : is_idle arr_seq sched t
exists j : Job, scheduled_at sched j t
exfalso ; eapply instant_t_is_not_idle in IDLE => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 IDLE : is_idle arr_seq sched t
busy_interval_prefix arr_seq sched j t1 t2
by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
Qed .
(** We then prove that if interference is [false] at a time [t]
then the job is scheduled. *)
Lemma not_interference_implies_scheduled :
~~ interference j t -> receives_service_at sched j t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2
~~ interference j t -> receives_service_at sched j t
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2
~~ interference j t -> receives_service_at sched j t
rewrite !negb_or /another_hep_job_interference => /andP [/andP [HYP1 HYP2] /hasPn HYP3].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP1 : ~~ is_blackout sched t HYP2 : ~~ service_inversion arr_seq sched j t HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j}
receives_service_at sched j t
move : HYP1; rewrite /is_blackout negbK => SUP; apply ideal_progress_inside_supplies => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP2 : ~~ service_inversion arr_seq sched j t HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t
scheduled_at sched j t
move : HYP2; rewrite negb_and negbK => /orP [SERV | /hasPn SI].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SERV : j \in served_jobs_at arr_seq sched t
scheduled_at sched j t
{ Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SERV : j \in served_jobs_at arr_seq sched t
scheduled_at sched j t
by apply service_at_implies_scheduled_at; apply : served_at_and_receives_service_consistent => //. } Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j}
scheduled_at sched j t
move : busy_implies_not_idle => [jo SCHED].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j} jo : Job SCHED : scheduled_at sched jo t
scheduled_at sched j t
have RSERV : receives_service_at sched jo t by apply ideal_progress_inside_supplies.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j} jo : Job SCHED : scheduled_at sched jo t RSERV : receives_service_at sched jo t
scheduled_at sched j t
have INRSERV : jo \in served_jobs_at arr_seq sched t by apply receives_service_and_served_at_consistent.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j} jo : Job SCHED : scheduled_at sched jo t RSERV : receives_service_at sched jo t INRSERV : jo \in served_jobs_at arr_seq sched t
scheduled_at sched j t
move : (HYP3 _ INRSERV); rewrite negb_and => /orP [LP | EQ]; last first .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j} jo : Job SCHED : scheduled_at sched jo t RSERV : receives_service_at sched jo t INRSERV : jo \in served_jobs_at arr_seq sched t EQ : ~~ (jo != j)
scheduled_at sched j t
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j} jo : Job SCHED : scheduled_at sched jo t RSERV : receives_service_at sched jo t INRSERV : jo \in served_jobs_at arr_seq sched t EQ : ~~ (jo != j)
scheduled_at sched j t
by move : EQ; rewrite negbK => /eqP EQ; subst jo.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 HYP3 : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ another_hep_job x j} SUP : has_supply sched t SI : {in served_jobs_at arr_seq sched t,
forall x : Job, ~~ ~~ hep_job_at t x j} jo : Job SCHED : scheduled_at sched jo t RSERV : receives_service_at sched jo t INRSERV : jo \in served_jobs_at arr_seq sched t LP : ~~ hep_job jo j
scheduled_at sched j t
by move : (SI _ INRSERV); rewrite LP.
Qed .
(** Conversely, if the job is scheduled at [t] then interference is [false]. *)
Lemma scheduled_implies_no_interference :
receives_service_at sched j t -> ~~ interference j t.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2
receives_service_at sched j t -> ~~ interference j t
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2
receives_service_at sched j t -> ~~ interference j t
move => RSERV.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 RSERV : receives_service_at sched j t
~~ interference j t
apply /negP => /orP [/orP[BL|PINV] | INT].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 RSERV : receives_service_at sched j t BL : is_blackout sched t
False
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 RSERV : receives_service_at sched j t BL : is_blackout sched t
False
by apply /negP; first apply : no_blackout_when_service_received.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 RSERV : receives_service_at sched j t PINV : service_inversion arr_seq sched j t
False
by apply /negP; first apply : receives_service_implies_no_service_inversion.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 RSERV : receives_service_at sched j t INT : another_hep_job_interference arr_seq sched j t
False
move : INT; rewrite_neg @no_ahep_interference_when_served.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job H_arrives : arrives_in arr_seq j H_job_cost_positive : 0 < job_cost jt1, t2 : instant H_busy_interval_prefix : busy_interval_prefix_ab j t1
t2 t : instant H_t_in_busy_interval : t1 <= t < t2 RSERV : receives_service_at sched j t
has_supply sched t
by apply : receives_service_implies_has_supply.
Qed .
End Abstract_Work_Conservation .
(** Using the above two lemmas, we can prove that abstract work
conservation always holds for these instantiations of
interference (I) and interfering workload (W). *)
Corollary instantiated_i_and_w_are_coherent_with_schedule :
work_conserving_ab.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl
work_conserving_ab
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl
work_conserving_ab
move => j t1 t2 t ARR POS BUSY NEQ; split .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job t1, t2 : instant t : nat ARR : arrives_in arr_seq j POS : 0 < job_cost jBUSY : definitions.busy_interval_prefix sched j t1 t2 NEQ : t1 <= t < t2
~ interference j t -> receives_service_at sched j t
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job t1, t2 : instant t : nat ARR : arrives_in arr_seq j POS : 0 < job_cost jBUSY : definitions.busy_interval_prefix sched j t1 t2 NEQ : t1 <= t < t2
~ interference j t -> receives_service_at sched j t
by move => G; apply : (not_interference_implies_scheduled j ARR POS); eauto 2 ; apply /negP.
- Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl j : Job t1, t2 : instant t : nat ARR : arrives_in arr_seq j POS : 0 < job_cost jBUSY : definitions.busy_interval_prefix sched j t1 t2 NEQ : t1 <= t < t2
receives_service_at sched j t -> ~ interference j t
by move => SERV; apply /negP; apply : scheduled_implies_no_interference; eauto 2 .
Qed .
(** Next, in order to prove that these definitions of interference
and interfering workload are consistent with sequential tasks,
we need to assume that the policy under consideration respects
sequential tasks. *)
Hypothesis H_policy_respects_sequential_tasks : policy_respects_sequential_tasks JLFP.
(** We prove that these definitions of interference and
interfering workload 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.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP
interference_and_workload_consistent_with_sequential_tasks
arr_seq sched tsk
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP
interference_and_workload_consistent_with_sequential_tasks
arr_seq sched tsk
move => j t1 t2 ARR /eqP TSK POS BUSY.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : definitions.busy_interval sched j t1 t2
task_workload_between arr_seq tsk 0 t1 =
task_service_of_jobs_in sched tsk
(arrivals_between arr_seq 0 t1) 0 t1
eapply instantiated_busy_interval_equivalent_busy_interval in BUSY => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2
task_workload_between arr_seq tsk 0 t1 =
task_service_of_jobs_in sched tsk
(arrivals_between arr_seq 0 t1) 0 t1
eapply all_jobs_have_completed_equiv_workload_eq_service => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2
forall j : Job,
j \in arrivals_between arr_seq 0 t1 ->
job_of_task tsk j -> completed_by sched j t1
move => s INs /eqP TSKs.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk
completed_by sched s t1
move : (INs) => NEQ.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk NEQ : s \in arrivals_between arr_seq 0 t1
completed_by sched s t1
eapply in_arrivals_implies_arrived_between in NEQ => //.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk NEQ : arrived_between s 0 t1
completed_by sched s t1
move : NEQ => /andP [_ JAs].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk JAs : job_arrival s < t1
completed_by sched s t1
move : (BUSY) => [[ _ [QT [_ /andP [JAj _]]] _]].Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk JAs : job_arrival s < t1 QT : quiet_time arr_seq sched j t1 JAj : t1 <= job_arrival j
completed_by sched s t1
apply QT => //; first exact : in_arrivals_implies_arrived.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk JAs : job_arrival s < t1 QT : quiet_time arr_seq sched j t1 JAj : t1 <= job_arrival j
hep_job s j
apply H_policy_respects_sequential_tasks; first by rewrite TSK TSKs.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1, t2 : instant ARR : arrives_in arr_seq j TSK : job_task j = tsk POS : 0 < job_cost jBUSY : busy_interval arr_seq sched j t1 t2 s : Job INs : s \in arrivals_between arr_seq 0 t1 TSKs : job_task s = tsk JAs : job_arrival s < t1 QT : quiet_time arr_seq sched j t1 JAj : t1 <= job_arrival j
job_arrival s <= job_arrival j
by apply leq_trans with t1; [lia |].
Qed .
(** Finally, we show that cumulative interference (I) never exceeds
cumulative interfering workload (W). *)
Lemma instantiated_i_and_w_no_speculative_execution :
no_speculative_execution.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP
no_speculative_execution
Proof .Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP
no_speculative_execution
move => j t1.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1 : nat
cumulative_interference j 0 t1 <=
cumulative_interfering_workload j 0 t1
rewrite cumulative_interference_split cumulative_interfering_workload_split.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1 : nat
blackout_during sched 0 t1 +
cumulative_service_inversion arr_seq sched j 0 t1 +
cumulative_another_hep_job_interference arr_seq sched
j 0 t1 <=
blackout_during sched 0 t1 +
cumulative_service_inversion arr_seq sched j 0 t1 +
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t1
rewrite leq_add2l cumulative_i_ohep_eq_service_of_ohep //=.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1 : nat
service_of_other_hep_jobs arr_seq sched j 0 t1 <=
cumulative_other_hep_jobs_interfering_workload arr_seq
j 0 t1
rewrite cumulative_iw_hep_eq_workload_of_ohep.Task : TaskType H : TaskCost Task Job : JobType H0 : JobTask Job Task H1 : JobArrival Job H2 : JobCost Job PState : ProcessorState Job H_uniprocessor_proc_model : uniprocessor_model PState H_unit_supply_proc_model : unit_supply_proc_model
PState H_consumed_supply_proc_model : fully_consuming_proc_model PState arr_seq : arrival_sequence Job H_valid_arrival_sequence : valid_arrival_sequence
arr_seq sched : schedule PState H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence
sched arr_seq H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute
sched H_completed_jobs_dont_execute : completed_jobs_dont_execute
sched JLFP : JLFP_policy Job H_priority_is_reflexive : reflexive_job_priorities
JLFP tsk : Task JobReady0 : JobReady Job PState H_work_bearing_readiness : work_bearing_readiness
arr_seq sched H_sched_valid : valid_schedule sched arr_seq work_conserving_ab := definitions.work_conserving
arr_seq sched : Prop work_conserving_cl := work_conserving arr_seq sched : Prop busy_interval_prefix_ab := definitions.busy_interval_prefix
sched : Job ->
instant -> instant -> Prop busy_interval_prefix_cl := busy_interval_prefix arr_seq
sched : Job ->
instant -> instant -> Prop H_work_conserving : work_conserving_cl H_policy_respects_sequential_tasks : policy_respects_sequential_tasks
JLFP j : Job t1 : nat
service_of_other_hep_jobs arr_seq sched j 0 t1 <=
workload_of_other_hep_jobs arr_seq j 0 t1
by apply service_of_jobs_le_workload => //.
Qed .
End I_IW_correctness .
End JLFPInstantiation .