Built with Alectryon, running Coq+SerAPI v8.19.0+0.19.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.
[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]
[Loading ML file coq-elpi.elpi ... done]
[Loading ML file zify_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.task. Require Export prosa.analysis.facts.interference. (** Throughout this file, we assume ideal uni-processor schedules. *) Require Import prosa.model.processor.ideal. Require Export prosa.analysis.facts.model.ideal.priority_inversion. (** * JLFP instantiation of Interference and Interfering Workload for ideal uni-processor. *) (** In this module we instantiate functions Interference and Interfering Workload for ideal uni-processor schedulers with an arbitrary JLFP-policy that satisfies the sequential-tasks hypothesis. We also prove equivalence of Interference and Interfering Workload to the more conventional notions of service or workload. *) 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 valid arrival sequence with consistent arrivals ... *) Variable arr_seq : arrival_sequence Job. Hypothesis H_valid_arrival_sequence : valid_arrival_sequence arr_seq. (** ... and any ideal uni-processor schedule of this arrival sequence ... *) Variable sched : schedule (ideal.processor_state Job). Hypothesis H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched arr_seq. (** ... 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. Hypothesis H_priority_is_transitive : transitive_job_priorities JLFP. (** We also assume that the policy respects sequential tasks, meaning that later-arrived jobs of a task don't have higher priority than earlier-arrived jobs of the same task. *) Hypothesis H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks JLFP. (** Let [tsk] be any task. *) Variable tsk : Task. (** * Interference and Interfering Workload *) (** In the following, we introduce definitions of interference, interfering workload and a function that bounds cumulative interference. *) (** ** Instantiation of Interference *) (** Before we define the notion of interference, we need to recall the definition of priority inversion. We say that job [j] is incurring a priority inversion at time [t] if there exists a job with lower priority that executes at time [t]. In order to simplify things, we ignore the fact that according to this definition a job can incur priority inversion even before its release (or after completion). All such (potentially bad) cases do not cause problems, as each job is analyzed only within the corresponding busy interval where the priority inversion behaves in the expected way. *) (** We say that job [j] incurs interference at time [t] iff it cannot execute due to a higher-or-equal-priority job being scheduled, or if it incurs a priority inversion. *) #[local,program] Instance ideal_jlfp_interference : Interference Job := { interference (j : Job) (t : instant) := priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t }. (** ** Instantiation of Interfering Workload *) (** The interfering workload, in turn, is defined as the sum of the priority inversion predicate and interfering workload of jobs with higher or equal priority. *) #[local,program] Instance ideal_jlfp_interfering_workload : InterferingWorkload Job := { interfering_workload (j : Job) (t : instant) := priority_inversion arr_seq sched j t + other_hep_jobs_interfering_workload arr_seq j t }. (** ** 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. 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 as well as a few useful rewrite rules. *) Section Equivalences. (** In the following subsection, we prove properties of the introduced functions under the assumption that the schedule is idle. *) Section IdleSchedule. (** Consider a time instant [t] ... *) Variable t : instant. (** ... and assume that the schedule is idle at [t]. *) Hypothesis H_idle : ideal_is_idle sched t. (** We prove that in this case: ... *) (** ... there is no interference, ... *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t

forall j : Job, ~~ interference j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t

forall j : Job, ~~ interference j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job

~ interference j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job
PI: priority_inversion arr_seq sched j t

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job
HEPI: another_hep_job_interference arr_seq sched j t
False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job
PI: priority_inversion arr_seq sched j t

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job
PI: priority_inversion arr_seq sched j t
j': Job
schj': scheduled_at sched j' t

False
exact: ideal_sched_implies_not_idle schj' H_idle.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job
HEPI: another_hep_job_interference arr_seq sched j t

False
by move: HEPI => /negPn; rewrite no_hep_job_interference_when_idle // is_idle_def //. Qed. (** ... as well as no interference for [tsk]. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t

forall j : Job, ~~ task_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t

forall j : Job, ~~ task_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
t: instant
H_idle: ideal_is_idle sched t
j: Job

~~ (nonself arr_seq sched j t && interference j t)
by rewrite negb_and; apply/orP; right; apply no_interference_when_idle. Qed. End IdleSchedule. (** Next, we prove properties of the introduced functions under the assumption that the scheduler is not idle. *) Section ScheduledJob. (** Consider a job [j] of task [tsk]. In this subsection, job [j] is deemed to be the main job with respect to which the functions are computed. *) Variable j : Job. Hypothesis H_j_tsk : job_of_task tsk j. (** Consider a time instant [t]. *) Variable t : instant. (** In the next subsection, we consider a case when a job [j'] from the same task (as job [j]) is scheduled. *) Section FromSameTask. (** Consider a job [j'] that comes from task [tsk] and is scheduled at time instant [t]. *) Variable j' : Job. Hypothesis H_j'_tsk : job_of_task tsk j'. Hypothesis H_j'_sched : scheduled_at sched j' t. (** Similarly, there is no task interference, since in order to incur the task interference, a job from a distinct task must be scheduled. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_tsk: job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t

~~ task_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_tsk: job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t

~~ task_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_tsk: job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t

~~ task_served_at arr_seq sched (job_task j) t -> False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_tsk: job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t

~~ job_of_task (job_task j) j' -> False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_tsk: job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t

tsk != tsk -> False
by rewrite eq_refl. Qed. End FromSameTask. (** In the next subsection, we consider a case when a job [j'] from a task other than [j]'s task is scheduled. *) Section FromDifferentTask. (** Consider a job [j'] that _does_ _not_ comes from task [tsk] and is scheduled at time instant [t]. *) Variable j' : Job. Hypothesis H_j'_not_tsk : ~~ job_of_task tsk j'. Hypothesis H_j'_sched : scheduled_at sched j' t. (** Hence, if we assume that [j'] has higher-or-equal priority, ... *) Hypothesis H_j'_hep : hep_job j' j. (** Moreover, in this case, task [tsk] also incurs interference. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

task_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

task_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

nonself arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j
interference j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

nonself arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

~~ task_scheduled_at arr_seq sched (job_task j) t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

~~ job_of_task (job_task j) j'
by move: H_j_tsk => /eqP -> .
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

interference j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

j' \in served_jobs_at arr_seq sched t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j
another_hep_job j' j
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

j' \in served_jobs_at arr_seq sched t
by apply scheduled_at_implies_in_served_at => //.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

another_hep_job j' j
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
H_j_tsk: job_of_task tsk j
t: instant
j': Job
H_j'_not_tsk: ~~ job_of_task tsk j'
H_j'_sched: scheduled_at sched j' t
H_j'_hep: hep_job j' j

~~ same_task j j'
exact: (diff_task tsk). Qed. End FromDifferentTask. End ScheduledJob. (** We prove that we can split cumulative interference into two parts: (1) cumulative priority inversion and (2) cumulative interference from jobs with higher or equal priority. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task

forall (j : Job) (t1 t2 : nat), cumulative_interference j t1 t2 = cumulative_priority_inversion arr_seq sched j t1 t2 + cumulative_another_hep_job_interference arr_seq sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task

forall (j : Job) (t1 t2 : nat), cumulative_interference j t1 t2 = cumulative_priority_inversion arr_seq sched j t1 t2 + cumulative_another_hep_job_interference arr_seq sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task

forall (j : Job) (t1 t2 : nat), \sum_(t1 <= t < t2) true && ideal_jlfp_interference j t = cumulative_priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2: nat

\sum_(t1 <= t < t2) ideal_jlfp_interference j t = \sum_(t1 <= i < t2) (priority_inversion arr_seq sched j i + another_hep_job_interference arr_seq sched j i)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2: nat

forall i : nat, true -> ideal_jlfp_interference j i <= priority_inversion arr_seq sched j i + another_hep_job_interference arr_seq sched j i
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2: nat
forall i : nat, true -> priority_inversion arr_seq sched j i + another_hep_job_interference arr_seq sched j i <= ideal_jlfp_interference j i
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2: nat

forall i : nat, true -> ideal_jlfp_interference j i <= priority_inversion arr_seq sched j i + another_hep_job_interference arr_seq sched j i
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat

priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t <= priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t
by destruct (priority_inversion _ _ _ _), (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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2: nat

forall i : nat, true -> priority_inversion arr_seq sched j i + another_hep_job_interference arr_seq sched j i <= ideal_jlfp_interference j i
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2: nat

forall i : nat, true -> priority_inversion arr_seq sched j i + another_hep_job_interference arr_seq sched j i <= ideal_jlfp_interference j i
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
IDLE: ideal_is_idle sched t

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
IDLE: ideal_is_idle sched t

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
IDLE: ideal_is_idle sched t

false + another_hep_job_interference arr_seq sched j t <= false || another_hep_job_interference arr_seq sched j t
by rewrite add0n.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
PRIO: hep_job s j = true

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
PRIO: hep_job s j = false
priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
PRIO: hep_job s j = true

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t
by erewrite sched_hep_implies_no_priority_inversion => //; rewrite add0n.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
PRIO: hep_job s j = false

priority_inversion arr_seq sched j t + another_hep_job_interference arr_seq sched j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
PRIO: hep_job s j = false

true + another_hep_job_interference arr_seq sched j t <= true || 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
j: Job
t1, t2, t: nat
s: Job
SCHED: scheduled_at sched s t
PRIO: hep_job s j = false

true + another_hep_job_interference arr_seq sched j t <= true
by rewrite add1n ltnS //= leqn0 eqb0; erewrite no_ahep_interference_when_scheduled_lp; last by erewrite PRIO. } Qed. (** Similarly, we prove that we can split cumulative interfering workload into two parts: (1) cumulative priority inversion and (2) cumulative interfering workload from jobs with higher or equal priority. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task

forall (j : Job) (t1 t2 : nat), cumulative_interfering_workload j t1 t2 = cumulative_priority_inversion arr_seq sched j t1 t2 + cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task

forall (j : Job) (t1 t2 : nat), cumulative_interfering_workload j t1 t2 = cumulative_priority_inversion arr_seq sched j t1 t2 + cumulative_other_hep_jobs_interfering_workload arr_seq j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task

forall (j : Job) (t1 t2 : nat), \sum_(t1 <= t < t2) interfering_workload j t = \sum_(t1 <= t < t2) priority_inversion arr_seq sched j t + \sum_(t1 <= t < t2) other_hep_jobs_interfering_workload arr_seq j t
by move => j t1 t2; rewrite -big_split //=. Qed. (** Let [j] be any job of task [tsk]. Then the cumulative task interference received by job [j] is bounded to the sum of the cumulative priority inversion of job [j] and the cumulative interference incurred by job [j] due to higher-or-equal priority jobs from other tasks. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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_task_interference arr_seq sched j t1 t2 <= cumulative_priority_inversion arr_seq sched j t1 t2 + cumulative_another_task_hep_job_interference arr_seq sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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_task_interference arr_seq sched j t1 t2 <= cumulative_priority_inversion arr_seq sched j t1 t2 + cumulative_another_task_hep_job_interference arr_seq sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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_task_interference arr_seq sched j t1 R <= cumulative_priority_inversion arr_seq sched j t1 R + cumulative_another_task_hep_job_interference arr_seq sched j t1 R
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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 arr_seq sched) j t <= cumulative_priority_inversion arr_seq sched j t1 R + cumulative_another_task_hep_job_interference arr_seq sched j t1 R
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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 arr_seq sched) j i <= \sum_(i <- index_iota t1 R | (i \in index_iota t1 R) && true) (priority_inversion arr_seq sched j i + another_task_hep_job_interference arr_seq sched j i)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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 arr_seq sched) j t <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
IDLE: ideal_is_idle sched t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
IDLE: ideal_is_idle sched t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
IIDLE: ideal_is_idle sched t
IDLE: is_idle arr_seq sched t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
IIDLE: ideal_is_idle sched t
IDLE: is_idle arr_seq sched t

~~ task_served_at arr_seq sched (job_task j) t && (false || another_hep_job_interference arr_seq sched j t) <= false + 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
IIDLE: ideal_is_idle sched t
IDLE: is_idle arr_seq sched t

~~ task_served_at arr_seq sched (job_task j) t && (false || false) <= false + 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
IIDLE: ideal_is_idle sched t
IDLE: is_idle arr_seq sched t

~~ task_served_at arr_seq sched (job_task j) t && (false || false) <= false + false
by rewrite andbF.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job_interference arr_seq sched j t) <= priority_inversion arr_seq sched j t + another_task_hep_job_interference arr_seq sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job s j) <= priority_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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ task_served_at arr_seq sched (job_task j) t && (priority_inversion arr_seq sched j t || another_hep_job s j) <= priority_inversion arr_seq sched j t + another_task_hep_job s j
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ task_served_at arr_seq sched (job_task j) t && (~~ hep_job s j || another_hep_job s j) <= ~~ hep_job s j + another_task_hep_job s j
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ job_of_task (job_task j) s && (~~ hep_job s j || another_hep_job s j) <= ~~ hep_job s j + another_task_hep_job s j
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && (s != j)) <= ~~ hep_job 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
EQj: s = j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ true) <= ~~ hep_job 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
EQj: s = j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ true) <= ~~ hep_job s j + hep_job s j && (job_task s != job_task j)
by subst; rewrite /job_of_task eq_refl H_priority_is_reflexive.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
EQt: job_task s == job_task j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job s j + hep_job s j && ~~ true
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
NEQt: job_task s != job_task j
~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job s j + hep_job s j && ~~ false
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
EQt: job_task s == job_task j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job s j + hep_job s j && ~~ true
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
EQt: job_task s == job_task j

~~ job_of_task (job_task s) s && (~~ hep_job s j || hep_job s j && ~~ false) - (~~ hep_job s j + hep_job s j && ~~ true) = 0
by rewrite /job_of_task eq_refl //= andbF addn0 eq_sym eqb0; apply/negP => LPs.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
NEQt: job_task s != job_task j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job s j + hep_job s j && ~~ false
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
s: Job
SCHEDs: scheduled_at sched s t
NEQj: s != j
NEQt: job_task s != job_task j

~~ job_of_task (job_task j) s && (~~ hep_job s j || hep_job s j && ~~ false) <= ~~ hep_job s j + hep_job s j && ~~ false
by rewrite /job_of_task NEQt //= andbT; case: hep_job. } } Qed. (** In this section, we prove that the (abstract) cumulative interfering workload is equivalent to the conventional workload, i.e., the one defined with concrete schedule parameters. *) 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. (** Then for any job [j], the cumulative interfering workload is equal to the conventional workload. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) (arrivals_between arr_seq t1 t2)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 t2 | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + 0) | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0
\sum_(t1 <= t < t1 + k.+1) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k.+1) | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + 0) | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 t1 | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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 = \sum_(j0 <- arrivals_between arr_seq t1 t1 | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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 = \sum_(j0 <- [::] | another_hep_job j0 j) job_cost j0
by rewrite /workload_of_jobs big_nil.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0

\sum_(t1 <= t < t1 + k.+1) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k.+1) | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0

\sum_(t1 <= i < t1 + k) \sum_(jhp <- arrivals_at arr_seq i | another_hep_job jhp j) job_cost jhp + \sum_(jhp <- arrivals_at arr_seq (t1 + k) | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k).+1 | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0

\sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0 + \sum_(jhp <- arrivals_at arr_seq (t1 + k) | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k).+1 | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0

\sum_(j0 <- \cat_(t1<=t<t1 + k)arrivals_at arr_seq t | another_hep_job j0 j) job_cost j0 + \sum_(jhp <- arrivals_at arr_seq (t1 + k) | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- (\cat_(t1<=i<t1 + k)arrivals_at arr_seq i ++ arrivals_at arr_seq (t1 + k)) | another_hep_job j0 j) job_cost j0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0
t1 <= t1 + k
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0

\sum_(j0 <- \cat_(t1<=t<t1 + k)arrivals_at arr_seq t | another_hep_job j0 j) job_cost j0 + \sum_(jhp <- arrivals_at arr_seq (t1 + k) | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- (\cat_(t1<=i<t1 + k)arrivals_at arr_seq i ++ arrivals_at arr_seq (t1 + k)) | another_hep_job j0 j) job_cost j0
by rewrite big_cat.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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) \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp = \sum_(j0 <- arrivals_between arr_seq t1 (t1 + k) | another_hep_job j0 j) job_cost j0

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. Hypothesis H_job_cost_positive : job_cost_positive j. (** To show the equivalence of the notions of busy intervals, we first show that the notions of quiet time are also equivalent. *) (** First, we show that the classical notion of quiet time implies the abstract notion of quiet time. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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 [H6 H7].
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall t : instant, 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall t : instant, 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive j

forall t : instant, 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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))
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive j
_t_: instant
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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_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
H_job_cost_positive: job_cost_positive 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 busy interval prefix obtained by instantiating the abstract definition of busy interval prefix coincides with the conventional definition of busy interval prefix. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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 t
REL: t1 <= job_arrival j < 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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 t
REL: 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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 t
REL: t1 <= job_arrival j < t2
forall t : nat, t1 < t < t2 -> ~ definitions.quiet_time sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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 t
REL: 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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 t
REL: 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall t1 t2 : instant, 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive j

forall t1 t2 : instant, 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks 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
H_job_cost_positive: job_cost_positive 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 (ideal.processor_state Job)}. 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. Let busy_interval_prefix_ab := definitions.busy_interval_prefix sched. (** We assume that the schedule is a work-conserving schedule in the _classical_ sense, and later prove that the hypothesis about abstract work-conservation also holds. *) Hypothesis H_work_conserving : work_conserving_cl. (** Assume the scheduling policy under consideration is reflexive. *) Hypothesis H_policy_reflexive : reflexive_job_priorities JLFP. (** 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 prove that if interference is [false] at a time [t] then the job is scheduled. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2

~~ priority_inversion arr_seq sched j t && ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t) -> receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
Idle: ideal_is_idle sched t

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
Idle: ideal_is_idle sched t

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
Idle: ideal_is_idle sched t

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix arr_seq sched j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
Idle: ideal_is_idle sched t

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix arr_seq sched j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
Idle: ideal_is_idle sched t

is_idle arr_seq sched t
by rewrite is_idle_def.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched j t

~~ another_hep_job jo j -> receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched j t
NHEP: ~~ hep_job jo j

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched j t
EQ: ~~ (jo != j)
receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched j t
NHEP: ~~ hep_job jo j

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched j t
NHEP: ~~ hep_job jo j
ZS: service_at sched j t = 0

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
NHEP: ~~ hep_job jo j
ZS: ~~ scheduled_at sched j t

priority_inversion arr_seq sched j t
exact/uni_priority_inversion_P.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
jo: Job
Sched_jo: scheduled_at sched jo t
PINV: ~ priority_inversion arr_seq sched j t
EQ: ~~ (jo != j)

receives_service_at sched j t
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
HYP1: ~~ priority_inversion arr_seq sched j t
HYP2: ~~ has (another_hep_job^~ j) (served_jobs_at arr_seq sched t)
Sched_jo: scheduled_at sched j t
PINV: ~ priority_inversion arr_seq sched j t

receives_service_at sched j t
by rewrite /receives_service_at service_at_is_scheduled_at Sched_jo. } Qed. (** Conversely, if the job is scheduled at [t] then interference is [false]. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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: priority_inversion arr_seq sched j t

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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: priority_inversion arr_seq sched j t

False
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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

scheduled_at sched j t
by rewrite /receives_service_at service_at_is_scheduled_at lt0b in RSERV.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
H_arrives: arrives_in arr_seq j
H_job_cost_positive: 0 < job_cost j
t1, t2: instant
H_busy_interval_prefix: busy_interval_prefix_ab j t1 t2
t: instant
H_t_in_busy_interval: t1 <= t < t2
INT: another_hep_job_interference arr_seq sched j t
RSERV: scheduled_at sched j t

False
by move: INT; rewrite_neg @no_ahep_interference_when_scheduled. Qed. End Abstract_Work_Conservation. (** Using the above two lemmas, we can prove that abstract work conservation always holds for these instantiations of [I] and [IW]. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP

work_conserving_ab
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP

work_conserving_ab
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
t1, t2: instant
t: nat
ARR: arrives_in arr_seq j
POS: 0 < job_cost j
BUSY: 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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
t1, t2: instant
t: nat
ARR: arrives_in arr_seq j
POS: 0 < job_cost j
BUSY: definitions.busy_interval_prefix sched j t1 t2
NEQ: t1 <= t < t2
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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
t1, t2: instant
t: nat
ARR: arrives_in arr_seq j
POS: 0 < job_cost j
BUSY: definitions.busy_interval_prefix sched j t1 t2
NEQ: t1 <= t < t2

~ interference j t -> receives_service_at sched j t
exact: (not_interference_implies_scheduled j ARR POS).
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
j: Job
t1, t2: instant
t: nat
ARR: arrives_in arr_seq j
POS: 0 < job_cost j
BUSY: definitions.busy_interval_prefix sched j t1 t2
NEQ: t1 <= t < t2

receives_service_at sched j t -> ~ interference j t
exact: scheduled_implies_no_interference. Qed. (** Next, in order to prove that these definitions of [I] and [IW] are consistent with sequential tasks, we need to assume that the policy under consideration respects sequential tasks. *) Hypothesis H_policy_respects_sequential_tasks : policy_respects_sequential_tasks JLFP. (** We prove that these definitions of [I] and [IW] are consistent with sequential tasks. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP

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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP

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
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
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 j
BUSY: 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. (** Since interfering and interfering workload are sufficient to define the busy window, next, we reason about the bound on the length of the busy window. *) Section BusyWindowBound. (** Consider an arrival curve. *) Context `{MaxArrivals Task}. (** Consider a set of tasks that respects the arrival curve. *) Variable ts : list Task. Hypothesis H_taskset_respects_max_arrivals : taskset_respects_max_arrivals arr_seq ts. (** Assume that all jobs come from this task set. *) Hypothesis H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts. (** Consider a constant [L] such that... *) Variable L : duration. (** ... [L] is greater than [0], and... *) Hypothesis H_L_positive : L > 0. (** [L] is the fixed point of the following equation. *) Hypothesis H_fixed_point : L = total_request_bound_function ts L. (** Assume all jobs have a valid job cost. *) Hypothesis H_arrivals_have_valid_job_costs : arrivals_have_valid_job_costs arr_seq. (** Then, we prove that [L] is a bound on the length of the busy window. *)
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq

busy_intervals_are_bounded_by arr_seq sched tsk L
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq

busy_intervals_are_bounded_by arr_seq sched tsk L
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j

exists t1 t2 : nat, t1 <= job_arrival j < t2 /\ t2 <= t1 + L /\ definitions.busy_interval sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j

forall t : instant, no_carry_in arr_seq sched t -> blackout_during sched t (t + L) + workload_of_jobs predT (arrivals_between arr_seq t (t + L)) <= L
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t1, t2: nat
T1: t1 <= job_arrival j < t2
T2: t2 <= t1 + L
GGG: busy_interval arr_seq sched j t1 t2
exists t1 t2 : nat, t1 <= job_arrival j < t2 /\ t2 <= t1 + L /\ definitions.busy_interval sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j

forall t : instant, no_carry_in arr_seq sched t -> blackout_during sched t (t + L) + workload_of_jobs predT (arrivals_between arr_seq t (t + L)) <= L
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t: instant

blackout_during sched t (t + L) + workload_of_jobs predT (arrivals_between arr_seq t (t + L)) <= total_request_bound_function ts L
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t: instant

blackout_during sched t (t + L) = 0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t: instant
0 + workload_of_jobs predT (arrivals_between arr_seq t (t + L)) <= total_request_bound_function ts L
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t: instant

blackout_during sched t (t + L) = 0
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t: instant
l: nat

is_blackout sched l = 0
by rewrite /is_blackout ideal_proc_has_supply.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t: instant

0 + workload_of_jobs predT (arrivals_between arr_seq t (t + L)) <= total_request_bound_function ts L
by rewrite add0n; apply total_workload_le_total_rbf.
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t1, t2: nat
T1: t1 <= job_arrival j < t2
T2: t2 <= t1 + L
GGG: busy_interval arr_seq sched j t1 t2

exists t1 t2 : nat, t1 <= job_arrival j < t2 /\ t2 <= t1 + L /\ definitions.busy_interval sched j t1 t2
Task: TaskType
H: TaskCost Task
Job: JobType
H0: JobTask Job Task
H1: JobArrival Job
H2: JobCost Job
arr_seq: arrival_sequence Job
H_valid_arrival_sequence: valid_arrival_sequence arr_seq
sched: schedule (processor_state Job)
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
H_priority_is_transitive: transitive_job_priorities JLFP
H_JLFP_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
tsk: Task
JobReady0: JobReady Job (processor_state Job)
H_work_bearing_readiness: work_bearing_readiness arr_seq sched
H_sched_valid: valid_schedule sched arr_seq
work_conserving_ab:= work_conserving arr_seq sched: Prop
work_conserving_cl:= work_conserving.work_conserving arr_seq sched: Prop
busy_interval_prefix_ab:= definitions.busy_interval_prefix sched: Job -> instant -> instant -> Prop
H_work_conserving: work_conserving_cl
H_policy_reflexive: reflexive_job_priorities JLFP
H_policy_respects_sequential_tasks: policy_respects_sequential_tasks JLFP
H3: MaxArrivals Task
ts: seq Task
H_taskset_respects_max_arrivals: taskset_respects_max_arrivals arr_seq ts
H_all_jobs_from_taskset: all_jobs_from_taskset arr_seq ts
L: duration
H_L_positive: 0 < L
H_fixed_point: L = total_request_bound_function ts L
H_arrivals_have_valid_job_costs: arrivals_have_valid_job_costs arr_seq
j: Job
ARR: arrives_in arr_seq j
TSK: job_of_task tsk j
POS: 0 < job_cost j
t1, t2: nat
T1: t1 <= job_arrival j < t2
T2: t2 <= t1 + L
GGG: busy_interval arr_seq sched j t1 t2

definitions.busy_interval sched j t1 t2
by apply instantiated_busy_interval_equivalent_busy_interval. Qed. End BusyWindowBound. End I_IW_correctness. End JLFPInstantiation. (** To preserve modularity and hide the implementation details of a technical definition presented in this file, we make the definition opaque. This way, we ensure that the system will treat each of these definitions as a single entity. *) Global Opaque another_hep_job_interference another_task_hep_job_interference ideal_jlfp_interference ideal_jlfp_interfering_workload cumulative_another_hep_job_interference cumulative_another_task_hep_job_interference cumulative_other_hep_jobs_interfering_workload other_hep_jobs_interfering_workload. (** We add some facts into the "Hint Database" basic_rt_facts, so Coq will be able to apply them automatically where needed. *) Global Hint Resolve abstract_busy_interval_classic_quiet_time abstract_busy_interval_classic_busy_interval_prefix : basic_rt_facts.